From 6d22793ddf002533d5d1abc3befe59ffc501a99f Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Mon, 6 May 2024 10:08:43 +0200 Subject: [PATCH 01/44] refactor: Array.feraseIdx: avoid have in definition (#4074) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit otherwise it remains in the equational theorem and may cause the “unused have linter” to trigger. By moving the proof into `decreasing_by`, the equational theorems are unencumbered by termination arguments. see also https://github.com/leanprover/std4/pull/690#issuecomment-2095378609 --- src/Init/Data/Array/Basic.lean | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/src/Init/Data/Array/Basic.lean b/src/Init/Data/Array/Basic.lean index 98993600555a..eacc93235ef6 100644 --- a/src/Init/Data/Array/Basic.lean +++ b/src/Init/Data/Array/Basic.lean @@ -733,17 +733,15 @@ def feraseIdx (a : Array α) (i : Fin a.size) : Array α := if h : i.val + 1 < a.size then let a' := a.swap ⟨i.val + 1, h⟩ i let i' : Fin a'.size := ⟨i.val + 1, by simp [a', h]⟩ - have : a'.size - i' < a.size - i := by - simp [a', Nat.sub_succ_lt_self _ _ i.isLt] a'.feraseIdx i' else a.pop termination_by a.size - i.val -decreasing_by simp_wf; decreasing_trivial_pre_omega +decreasing_by simp_wf; exact Nat.sub_succ_lt_self _ _ i.isLt theorem size_feraseIdx (a : Array α) (i : Fin a.size) : (a.feraseIdx i).size = a.size - 1 := by induction a, i using Array.feraseIdx.induct with - | @case1 a i h a' _ _ ih => + | @case1 a i h a' _ ih => unfold feraseIdx simp [h, a', ih] | case2 a i h => From 9e4c414f48eb644edb0d32c6c77e0aea4e694a7d Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Mon, 6 May 2024 04:08:19 -0700 Subject: [PATCH 02/44] fix: remove `Subtype.instInhabited` (#4055) It has had a long history going back [10 years](https://github.com/leanprover-community/lean/commit/3afad10a7294fd183f66cc9d4723511b799ac4ff#diff-4e22e2bb74f004d2ff7cdabcb5c01429abbc906e20befe2517679e257b4387e4R41), but its time has come to an end since this instance is never applicable. --- src/Init/Core.lean | 3 --- 1 file changed, 3 deletions(-) diff --git a/src/Init/Core.lean b/src/Init/Core.lean index 89d801ec98d0..144224ac5401 100644 --- a/src/Init/Core.lean +++ b/src/Init/Core.lean @@ -1114,9 +1114,6 @@ theorem eta (a : {x // p x}) (h : p (val a)) : mk (val a) h = a := by cases a exact rfl -instance {α : Type u} {p : α → Prop} {a : α} (h : p a) : Inhabited {x // p x} where - default := ⟨a, h⟩ - instance {α : Type u} {p : α → Prop} [DecidableEq α] : DecidableEq {x : α // p x} := fun ⟨a, h₁⟩ ⟨b, h₂⟩ => if h : a = b then isTrue (by subst h; exact rfl) From 35d9307df30c1cbd2ab1103e083c21dd70dbd191 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 6 May 2024 21:14:18 +1000 Subject: [PATCH 03/44] chore: move @[simp] attribute on length_eq_zero earlier (#4077) Cleanup. --- src/Init/Data/List/Lemmas.lean | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index 702ebddcb2d1..00070de044fb 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -54,7 +54,7 @@ theorem eq_nil_of_length_eq_zero (_ : length l = 0) : l = [] := match l with | [ theorem ne_nil_of_length_eq_succ (_ : length l = succ n) : l ≠ [] := fun _ => nomatch l -theorem length_eq_zero : length l = 0 ↔ l = [] := +@[simp] theorem length_eq_zero : length l = 0 ↔ l = [] := ⟨eq_nil_of_length_eq_zero, fun h => h ▸ rfl⟩ /-! ### mem -/ @@ -784,8 +784,6 @@ theorem exists_cons_of_length_succ : ∀ {l : List α}, l.length = n + 1 → ∃ h t, l = h :: t | _::_, _ => ⟨_, _, rfl⟩ -attribute [simp] length_eq_zero -- TODO: suggest to core - @[simp] theorem length_pos {l : List α} : 0 < length l ↔ l ≠ [] := Nat.pos_iff_ne_zero.trans (not_congr length_eq_zero) From 00dceb9a9d92b83dafcf59179f399dcf5ce5b5e4 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 May 2024 21:17:35 +0200 Subject: [PATCH 04/44] fix: code duplication at `liftCoreM` and `liftTermElabM` at `Command.lean` (#4080) This PR also fixes: - Fields caching specific `Options` at `CoreM` are now properly set. - `nextMacroScope` was not being propagated at `liftCoreM`. --- src/Lean/CoreM.lean | 8 +++- src/Lean/Elab/Command.lean | 76 ++++++++++++++++++-------------------- 2 files changed, 41 insertions(+), 43 deletions(-) diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index 76f9c36315f7..b2d76c7b3245 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -131,6 +131,10 @@ instance : MonadWithOptions CoreM where maxRecDepth := maxRecDepth.get options }) x +-- Helper function for ensuring fields that depend on `options` have the correct value. +@[inline] private def withConsistentCtx (x : CoreM α) : CoreM α := do + withOptions id x + instance : AddMessageContext CoreM where addMessageContext := addMessageContextPartial @@ -217,7 +221,7 @@ def mkFreshUserName (n : Name) : CoreM Name := mkFreshNameImp n @[inline] def CoreM.run (x : CoreM α) (ctx : Context) (s : State) : EIO Exception (α × State) := - (x ctx).run s + ((withConsistentCtx x) ctx).run s @[inline] def CoreM.run' (x : CoreM α) (ctx : Context) (s : State) : EIO Exception α := Prod.fst <$> x.run ctx s @@ -231,7 +235,7 @@ def mkFreshUserName (n : Name) : CoreM Name := instance [MetaEval α] : MetaEval (CoreM α) where eval env opts x _ := do let x : CoreM α := do try x finally printTraces - let (a, s) ← (withOptions (fun _ => opts) x).toIO { fileName := "", fileMap := default } { env := env } + let (a, s) ← (withConsistentCtx x).toIO { fileName := "", fileMap := default, options := opts } { env := env } MetaEval.eval s.env opts a (hideUnit := true) -- withIncRecDepth for a monad `m` such that `[MonadControlT CoreM n]` diff --git a/src/Lean/Elab/Command.lean b/src/Lean/Elab/Command.lean index e2bf601aaeee..405faa4ef986 100644 --- a/src/Lean/Elab/Command.lean +++ b/src/Lean/Elab/Command.lean @@ -128,20 +128,6 @@ def mkMessageAux (ctx : Context) (ref : Syntax) (msgData : MessageData) (severit let endPos := ref.getTailPos?.getD pos mkMessageCore ctx.fileName ctx.fileMap msgData severity pos endPos -private def mkCoreContext (ctx : Context) (s : State) (heartbeats : Nat) : Core.Context := - let scope := s.scopes.head! - { fileName := ctx.fileName - fileMap := ctx.fileMap - options := scope.opts - currRecDepth := ctx.currRecDepth - maxRecDepth := s.maxRecDepth - ref := ctx.ref - currNamespace := scope.currNamespace - openDecls := scope.openDecls - initHeartbeats := heartbeats - currMacroScope := ctx.currMacroScope - diag := getDiag scope.opts } - private def addTraceAsMessagesCore (ctx : Context) (log : MessageLog) (traceState : TraceState) : MessageLog := Id.run do if traceState.traces.isEmpty then return log let mut traces : HashMap (String.Pos × String.Pos) (Array MessageData) := ∅ @@ -167,31 +153,49 @@ private def addTraceAsMessages : CommandElabM Unit := do traceState.traces := {} } -def liftCoreM (x : CoreM α) : CommandElabM α := do +private def runCore (x : CoreM α) : CommandElabM α := do let s ← get let ctx ← read let heartbeats ← IO.getNumHeartbeats - let Eα := Except Exception α - let x : CoreM Eα := try let a ← x; pure <| Except.ok a catch ex => pure <| Except.error ex - let x : EIO Exception (Eα × Core.State) := (ReaderT.run x (mkCoreContext ctx s heartbeats)).run { env := s.env, ngen := s.ngen, traceState := s.traceState, messages := {}, infoState.enabled := s.infoState.enabled } + let env := s.env + let scope := s.scopes.head! + let coreCtx : Core.Context := { + fileName := ctx.fileName + fileMap := ctx.fileMap + currRecDepth := ctx.currRecDepth + maxRecDepth := s.maxRecDepth + ref := ctx.ref + currNamespace := scope.currNamespace + openDecls := scope.openDecls + initHeartbeats := heartbeats + currMacroScope := ctx.currMacroScope + options := scope.opts + } + let x : EIO _ _ := x.run coreCtx { + env + ngen := s.ngen + nextMacroScope := s.nextMacroScope + infoState.enabled := s.infoState.enabled + traceState := s.traceState + } let (ea, coreS) ← liftM x modify fun s => { s with - env := coreS.env - ngen := coreS.ngen - messages := s.messages ++ coreS.messages + env := coreS.env + nextMacroScope := coreS.nextMacroScope + ngen := coreS.ngen + infoState.trees := s.infoState.trees.append coreS.infoState.trees traceState.traces := coreS.traceState.traces.map fun t => { t with ref := replaceRef t.ref ctx.ref } - infoState.trees := s.infoState.trees.append coreS.infoState.trees + messages := s.messages ++ coreS.messages } - match ea with - | Except.ok a => pure a - | Except.error e => throw e + return ea + +def liftCoreM (x : CoreM α) : CommandElabM α := do + MonadExcept.ofExcept (← runCore (observing x)) private def ioErrorToMessage (ctx : Context) (ref : Syntax) (err : IO.Error) : Message := let ref := getBetterRef ref ctx.macroStack mkMessageAux ctx ref (toString err) MessageSeverity.error -@[inline] def liftEIO {α} (x : EIO Exception α) : CommandElabM α := liftM x - @[inline] def liftIO {α} (x : IO α) : CommandElabM α := do let ctx ← read IO.toEIO (fun (ex : IO.Error) => Exception.error ctx.ref ex.toString) x @@ -271,7 +275,7 @@ private def elabCommandUsing (s : State) (stx : Syntax) : List (KeyedDeclsAttrib (fun _ => do set s; elabCommandUsing s stx elabFns) /-- Elaborate `x` with `stx` on the macro stack -/ -def withMacroExpansion {α} (beforeStx afterStx : Syntax) (x : CommandElabM α) : CommandElabM α := +def withMacroExpansion (beforeStx afterStx : Syntax) (x : CommandElabM α) : CommandElabM α := withInfoContext (mkInfo := pure <| .ofMacroExpansionInfo { stx := beforeStx, output := afterStx, lctx := .empty }) do withReader (fun ctx => { ctx with macroStack := { before := beforeStx, after := afterStx } :: ctx.macroStack }) x @@ -404,7 +408,6 @@ def printExpr (e : Expr) : MetaM Unit := do def liftTermElabM (x : TermElabM α) : CommandElabM α := do let ctx ← read let s ← get - let heartbeats ← IO.getNumHeartbeats -- dbg_trace "heartbeats: {heartbeats}" let scope := s.scopes.head! -- We execute `x` with an empty message log. Thus, `x` cannot modify/view messages produced by previous commands. @@ -413,18 +416,9 @@ def liftTermElabM (x : TermElabM α) : CommandElabM α := do -- make sure `observing` below also catches runtime exceptions (like we do by default in -- `CommandElabM`) let _ := MonadAlwaysExcept.except (m := TermElabM) - let x : MetaM _ := (observing (try x finally Meta.reportDiag)).run (mkTermContext ctx s) { levelNames := scope.levelNames } - let x : CoreM _ := x.run mkMetaContext {} - let x : EIO _ _ := x.run (mkCoreContext ctx s heartbeats) { env := s.env, ngen := s.ngen, nextMacroScope := s.nextMacroScope, infoState.enabled := s.infoState.enabled, traceState := s.traceState } - let (((ea, _), _), coreS) ← liftEIO x - modify fun s => { s with - env := coreS.env - nextMacroScope := coreS.nextMacroScope - ngen := coreS.ngen - infoState.trees := s.infoState.trees.append coreS.infoState.trees - traceState.traces := coreS.traceState.traces.map fun t => { t with ref := replaceRef t.ref ctx.ref } - messages := s.messages ++ coreS.messages - } + let x : MetaM _ := (observing (try x finally Meta.reportDiag)).run (mkTermContext ctx s) { levelNames := scope.levelNames } + let x : CoreM _ := x.run mkMetaContext {} + let ((ea, _), _) ← runCore x MonadExcept.ofExcept ea /-- From 07c407ab8212e2d7ddcb24caa492f6943e7b7c92 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 May 2024 23:53:16 +0200 Subject: [PATCH 05/44] feat: collect kernel diagnostic information (#4082) We now also track which declarations have been unfolded by the kernel when using ```lean set_option diagnostics true ``` --- src/Lean/CoreM.lean | 9 ++-- src/Lean/Elab/Command.lean | 2 +- src/Lean/Environment.lean | 49 ++++++++++++++++++++ src/Lean/Meta/Diagnostics.lean | 2 + src/kernel/environment.cpp | 82 ++++++++++++++++++++++++++-------- src/kernel/environment.h | 33 ++++++++++++++ src/kernel/inductive.cpp | 14 +++--- src/kernel/type_checker.cpp | 19 +++++--- src/kernel/type_checker.h | 5 ++- tests/lean/run/ack.lean | 29 ++++++++++++ 10 files changed, 208 insertions(+), 36 deletions(-) diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index b2d76c7b3245..5664bd33196f 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -121,13 +121,16 @@ instance : MonadOptions CoreM where getOptions := return (← read).options instance : MonadWithOptions CoreM where - withOptions f x := + withOptions f x := do + let options := f (← read).options + let diag := diagnostics.get options + if Kernel.isDiagnosticsEnabled (← getEnv) != diag then + modifyEnv fun env => Kernel.enableDiag env diag withReader (fun ctx => - let options := f ctx.options { ctx with options - diag := diagnostics.get options + diag maxRecDepth := maxRecDepth.get options }) x diff --git a/src/Lean/Elab/Command.lean b/src/Lean/Elab/Command.lean index 405faa4ef986..590b139aea16 100644 --- a/src/Lean/Elab/Command.lean +++ b/src/Lean/Elab/Command.lean @@ -157,7 +157,7 @@ private def runCore (x : CoreM α) : CommandElabM α := do let s ← get let ctx ← read let heartbeats ← IO.getNumHeartbeats - let env := s.env + let env := Kernel.resetDiag s.env let scope := s.scopes.head! let coreCtx : Core.Context := { fileName := ctx.fileName diff --git a/src/Lean/Environment.lean b/src/Lean/Environment.lean index f8b736c03d91..1bb99cd02053 100644 --- a/src/Lean/Environment.lean +++ b/src/Lean/Environment.lean @@ -901,6 +901,55 @@ builtin_initialize namespacesExt : SimplePersistentEnvExtension Name NameSSet addEntryFn := fun s n => s.insert n } +structure Kernel.Diagnostics where + /-- Number of times each declaration has been unfolded by the kernel. -/ + unfoldCounter : PHashMap Name Nat := {} + /-- If `enabled = true`, kernel records declarations that have been unfolded. -/ + enabled : Bool := false + deriving Inhabited + +/-- +Extension for storting diagnostic information. + +Remark: We store kernel diagnostic information in an environment extension to simplify +the interface with the kernel implemented in C/C++. Thus, we can only track +declarations in methods, such as `addDecl`, which return a new environment. +`Kernel.isDefEq` and `Kernel.whnf` do not update the statistics. We claim +this is ok since these methods are mainly used for debugging. +-/ +builtin_initialize diagExt : EnvExtension Kernel.Diagnostics ← + registerEnvExtension (pure {}) + +@[export lean_kernel_diag_is_enabled] +def Kernel.Diagnostics.isEnabled (d : Diagnostics) : Bool := + d.enabled + +/-- Enables/disables kernel diagnostics. -/ +def Kernel.enableDiag (env : Environment) (flag : Bool) : Environment := + diagExt.modifyState env fun s => { s with enabled := flag } + +def Kernel.isDiagnosticsEnabled (env : Environment) : Bool := + diagExt.getState env |>.enabled + +def Kernel.resetDiag (env : Environment) : Environment := + diagExt.modifyState env fun s => { s with unfoldCounter := {} } + +@[export lean_kernel_record_unfold] +def Kernel.Diagnostics.recordUnfold (d : Diagnostics) (declName : Name) : Diagnostics := + if d.enabled then + let cNew := if let some c := d.unfoldCounter.find? declName then c + 1 else 1 + { d with unfoldCounter := d.unfoldCounter.insert declName cNew } + else + d + +@[export lean_kernel_get_diag] +def Kernel.getDiagnostics (env : Environment) : Diagnostics := + diagExt.getState env + +@[export lean_kernel_set_diag] +def Kernel.setDiagnostics (env : Environment) (diag : Diagnostics) : Environment := + diagExt.setState env diag + namespace Environment /-- Register a new namespace in the environment. -/ diff --git a/src/Lean/Meta/Diagnostics.lean b/src/Lean/Meta/Diagnostics.lean index 1bf034fe4296..8dd0537aa25f 100644 --- a/src/Lean/Meta/Diagnostics.lean +++ b/src/Lean/Meta/Diagnostics.lean @@ -75,6 +75,7 @@ def reportDiag : MetaM Unit := do let unfoldReducible ← mkDiagSummaryForUnfoldedReducible unfoldCounter let heu ← mkDiagSummary (← get).diag.heuristicCounter let inst ← mkDiagSummaryForUsedInstances + let unfoldKernel ← mkDiagSummary (Kernel.getDiagnostics (← getEnv)).unfoldCounter unless unfoldDefault.isEmpty && unfoldInstance.isEmpty && unfoldReducible.isEmpty && heu.isEmpty && inst.isEmpty do let m := MessageData.nil let m := appendSection m `reduction "unfolded declarations" unfoldDefault @@ -82,6 +83,7 @@ def reportDiag : MetaM Unit := do let m := appendSection m `reduction "unfolded reducible declarations" unfoldReducible let m := appendSection m `type_class "used instances" inst let m := appendSection m `def_eq "heuristic for solving `f a =?= f b`" heu + let m := appendSection m `kernel "unfolded declarations" unfoldKernel let m := m ++ "use `set_option diagnostics.threshold ` to control threshold for reporting counters" logInfo m diff --git a/src/kernel/environment.cpp b/src/kernel/environment.cpp index 792f4880dfe3..d8316d5c9f3e 100644 --- a/src/kernel/environment.cpp +++ b/src/kernel/environment.cpp @@ -28,6 +28,38 @@ extern "C" object* lean_get_extension(object*, object*); extern "C" object* lean_set_extension(object*, object*, object*); extern "C" object* lean_environment_set_main_module(object*, object*); extern "C" object* lean_environment_main_module(object*); +extern "C" object* lean_kernel_record_unfold (object*, object*); +extern "C" object* lean_kernel_get_diag(object*); +extern "C" object* lean_kernel_set_diag(object*, object*); +extern "C" uint8* lean_kernel_diag_is_enabled(object*); + +void diagnostics::record_unfold(name const & decl_name) { + m_obj = lean_kernel_record_unfold(to_obj_arg(), decl_name.to_obj_arg()); +} + +scoped_diagnostics::scoped_diagnostics(environment const & env, bool collect) { + if (collect) { + diagnostics d(env.get_diag()); + if (lean_kernel_diag_is_enabled(d.to_obj_arg())) { + m_diag = new diagnostics(d); + } else + m_diag = nullptr; + } else { + m_diag = nullptr; + } +} + +scoped_diagnostics::~scoped_diagnostics() { + if (m_diag) + delete m_diag; +} + +environment scoped_diagnostics::update(environment const & env) const { + if (m_diag) + return env.set_diag(*m_diag); + else + return env; +} environment mk_empty_environment(uint32 trust_lvl) { return get_io_result(lean_mk_empty_environment(trust_lvl, io_mk_world())); @@ -37,6 +69,14 @@ environment::environment(unsigned trust_lvl): object_ref(mk_empty_environment(trust_lvl)) { } +diagnostics environment::get_diag() const { + return diagnostics(lean_kernel_get_diag(to_obj_arg())); +} + +environment environment::set_diag(diagnostics const & diag) const { + return environment(lean_kernel_set_diag(to_obj_arg(), diag.to_obj_arg())); +} + void environment::set_main_module(name const & n) { m_obj = lean_environment_set_main_module(m_obj, n.to_obj_arg()); } @@ -118,13 +158,13 @@ static void check_constant_val(environment const & env, constant_val const & v, checker.ensure_sort(sort, v.get_type()); } -static void check_constant_val(environment const & env, constant_val const & v, definition_safety ds) { - type_checker checker(env, ds); +static void check_constant_val(environment const & env, constant_val const & v, diagnostics * diag, definition_safety ds) { + type_checker checker(env, diag, ds); check_constant_val(env, v, checker); } -static void check_constant_val(environment const & env, constant_val const & v, bool safe_only) { - check_constant_val(env, v, safe_only ? definition_safety::safe : definition_safety::unsafe); +static void check_constant_val(environment const & env, constant_val const & v, diagnostics * diag, bool safe_only) { + check_constant_val(env, v, diag, safe_only ? definition_safety::safe : definition_safety::unsafe); } void environment::add_core(constant_info const & info) { @@ -136,48 +176,50 @@ environment environment::add(constant_info const & info) const { } environment environment::add_axiom(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); axiom_val const & v = d.to_axiom_val(); if (check) - check_constant_val(*this, v.to_constant_val(), !d.is_unsafe()); - return add(constant_info(d)); + check_constant_val(*this, v.to_constant_val(), diag.get(), !d.is_unsafe()); + return diag.update(add(constant_info(d))); } environment environment::add_definition(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); definition_val const & v = d.to_definition_val(); if (v.is_unsafe()) { /* Meta definition can be recursive. So, we check the header, add, and then type check the body. */ if (check) { - type_checker checker(*this, definition_safety::unsafe); + type_checker checker(*this, diag.get(), definition_safety::unsafe); check_constant_val(*this, v.to_constant_val(), checker); } environment new_env = add(constant_info(d)); if (check) { - type_checker checker(new_env, definition_safety::unsafe); + type_checker checker(new_env, diag.get(), definition_safety::unsafe); check_no_metavar_no_fvar(new_env, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(new_env, d, val_type); } - return new_env; + return diag.update(new_env); } else { if (check) { - type_checker checker(*this); + type_checker checker(*this, diag.get()); 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()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } } environment environment::add_theorem(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); theorem_val const & v = d.to_theorem_val(); if (check) { - // TODO(Leo): we must add support for handling tasks here - type_checker checker(*this); + type_checker checker(*this, diag.get()); 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); @@ -186,22 +228,24 @@ environment environment::add_theorem(declaration const & d, bool check) const { if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } environment environment::add_opaque(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); opaque_val const & v = d.to_opaque_val(); if (check) { - type_checker checker(*this); + type_checker checker(*this, diag.get()); check_constant_val(*this, v.to_constant_val(), checker); expr val_type = checker.check(v.get_value(), v.get_lparams()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } environment environment::add_mutual(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); definition_vals const & vs = d.to_definition_vals(); if (empty(vs)) throw kernel_exception(*this, "invalid empty mutual definition"); @@ -210,7 +254,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { throw kernel_exception(*this, "invalid mutual definition, declaration is not tagged as unsafe/partial"); /* Check declarations header */ if (check) { - type_checker checker(*this, safety); + type_checker checker(*this, diag.get(), safety); for (definition_val const & v : vs) { if (v.get_safety() != safety) throw kernel_exception(*this, "invalid mutual definition, declarations must have the same safety annotation"); @@ -224,7 +268,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { } /* Check actual definitions */ if (check) { - type_checker checker(new_env, safety); + type_checker checker(new_env, diag.get(), safety); for (definition_val const & v : vs) { check_no_metavar_no_fvar(new_env, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); @@ -232,7 +276,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { throw definition_type_mismatch_exception(new_env, d, val_type); } } - return new_env; + return diag.update(new_env); } environment environment::add(declaration const & d, bool check) const { diff --git a/src/kernel/environment.h b/src/kernel/environment.h index d9cb69f47258..eaa077dbf16a 100644 --- a/src/kernel/environment.h +++ b/src/kernel/environment.h @@ -29,6 +29,36 @@ class environment_extension { virtual ~environment_extension() {} }; +/* Wrapper for `Kernel.Diagnostics` */ +class diagnostics : public object_ref { +public: + diagnostics(diagnostics const & other):object_ref(other) {} + diagnostics(diagnostics && other):object_ref(other) {} + explicit diagnostics(b_obj_arg o, bool b):object_ref(o, b) {} + explicit diagnostics(obj_arg o):object_ref(o) {} + ~diagnostics() {} + void record_unfold(name const & decl_name); +}; + +/* +Store `Kernel.Diagnostics` stored in environment extension in `m_diag` IF +- `Kernel.Diagnostics.enable = true` +- `collect = true`. This is a minor optimization. + +We use this class to ensure we don't waste time collecting information +that was not requested. +*/ +class scoped_diagnostics { + diagnostics * m_diag; +public: + scoped_diagnostics(environment const & env, bool collect); + scoped_diagnostics(scoped_diagnostics const &) = delete; + scoped_diagnostics(scoped_diagnostics &&) = delete; + ~scoped_diagnostics(); + environment update(environment const &) const; + diagnostics * get() const { return m_diag; } +}; + class environment : public object_ref { friend class add_inductive_fn; @@ -56,6 +86,9 @@ class environment : public object_ref { environment & operator=(environment const & other) { object_ref::operator=(other); return *this; } environment & operator=(environment && other) { object_ref::operator=(other); return *this; } + diagnostics get_diag() const; + environment set_diag(diagnostics const & diag) const; + /** \brief Return the trust level of this environment. */ unsigned trust_lvl() const; diff --git a/src/kernel/inductive.cpp b/src/kernel/inductive.cpp index 2238b24a256c..8b75cdeff2a7 100644 --- a/src/kernel/inductive.cpp +++ b/src/kernel/inductive.cpp @@ -124,6 +124,7 @@ optional get_rec_rule_for(recursor_val const & rec_val, expr cons class add_inductive_fn { environment m_env; name_generator m_ngen; + diagnostics * m_diag; local_ctx m_lctx; names m_lparams; unsigned m_nparams; @@ -158,8 +159,8 @@ class add_inductive_fn { buffer m_rec_infos; public: - add_inductive_fn(environment const & env, inductive_decl const & decl, bool is_nested): - m_env(env), m_ngen(*g_ind_fresh), m_lparams(decl.get_lparams()), m_is_unsafe(decl.is_unsafe()), + add_inductive_fn(environment const & env, diagnostics * diag, inductive_decl const & decl, bool is_nested): + m_env(env), m_ngen(*g_ind_fresh), m_diag(diag), m_lparams(decl.get_lparams()), m_is_unsafe(decl.is_unsafe()), m_is_nested(is_nested) { if (!decl.get_nparams().is_small()) throw kernel_exception(env, "invalid inductive datatype, number of parameters is too big"); @@ -167,7 +168,7 @@ class add_inductive_fn { to_buffer(decl.get_types(), m_ind_types); } - type_checker tc() { return type_checker(m_env, m_lctx, m_is_unsafe ? definition_safety::unsafe : definition_safety::safe); } + type_checker tc() { return type_checker(m_env, m_lctx, m_diag, m_is_unsafe ? definition_safety::unsafe : definition_safety::safe); } /** Return type of the parameter at position `i` */ expr get_param_type(unsigned i) const { @@ -1112,10 +1113,11 @@ static pair> mk_aux_rec_name_map(environment const & aux_e environment environment::add_inductive(declaration const & d) const { elim_nested_inductive_result res = elim_nested_inductive_fn(*this, d)(); bool is_nested = !res.m_aux2nested.empty(); - environment aux_env = add_inductive_fn(*this, inductive_decl(res.m_aux_decl), is_nested)(); + scoped_diagnostics diag(*this, true); + environment aux_env = add_inductive_fn(*this, diag.get(), inductive_decl(res.m_aux_decl), is_nested)(); if (!is_nested) { /* `d` did not contain nested inductive types. */ - return aux_env; + return diag.update(aux_env); } else { /* Restore nested inductives. */ inductive_decl ind_d(d); @@ -1170,7 +1172,7 @@ environment environment::add_inductive(declaration const & d) const { for (name const & aux_rec : aux_rec_names) { process_rec(aux_rec); } - return new_env; + return diag.update(new_env); } } diff --git a/src/kernel/type_checker.cpp b/src/kernel/type_checker.cpp index 9adb025818c2..1dbb0819a98e 100644 --- a/src/kernel/type_checker.cpp +++ b/src/kernel/type_checker.cpp @@ -477,6 +477,11 @@ expr type_checker::whnf_core(expr const & e, bool cheap_rec, bool cheap_proj) { cheap_rec, cheap_proj); } else if (f == f0) { if (auto r = reduce_recursor(e, cheap_rec, cheap_proj)) { + if (m_diag) { + auto f = get_app_fn(e); + if (is_constant(f)) + m_diag->record_unfold(const_name(f)); + } /* iota-reduction and quotient reduction rules */ return whnf_core(*r, cheap_rec, cheap_proj); } else { @@ -513,8 +518,12 @@ optional type_checker::is_delta(expr const & e) const { optional type_checker::unfold_definition_core(expr const & e) { if (is_constant(e)) { if (auto d = is_delta(e)) { - if (length(const_levels(e)) == d->get_num_lparams()) + if (length(const_levels(e)) == d->get_num_lparams()) { + if (m_diag) { + m_diag->record_unfold(d->get_name()); + } return some_expr(instantiate_value_lparams(*d, const_levels(e))); + } } } return none_expr(); @@ -1148,18 +1157,18 @@ expr type_checker::eta_expand(expr const & e) { return m_lctx.mk_lambda(fvars, r); } -type_checker::type_checker(environment const & env, local_ctx const & lctx, definition_safety ds): - m_st_owner(true), m_st(new state(env)), +type_checker::type_checker(environment const & env, local_ctx const & lctx, diagnostics * diag, definition_safety ds): + m_st_owner(true), m_st(new state(env)), m_diag(diag), m_lctx(lctx), m_definition_safety(ds), m_lparams(nullptr) { } type_checker::type_checker(state & st, local_ctx const & lctx, definition_safety ds): - m_st_owner(false), m_st(&st), m_lctx(lctx), + m_st_owner(false), m_st(&st), m_diag(nullptr), m_lctx(lctx), m_definition_safety(ds), m_lparams(nullptr) { } type_checker::type_checker(type_checker && src): - m_st_owner(src.m_st_owner), m_st(src.m_st), m_lctx(std::move(src.m_lctx)), + m_st_owner(src.m_st_owner), m_st(src.m_st), m_diag(src.m_diag), m_lctx(std::move(src.m_lctx)), m_definition_safety(src.m_definition_safety), m_lparams(src.m_lparams) { src.m_st_owner = false; } diff --git a/src/kernel/type_checker.h b/src/kernel/type_checker.h index e38a60772bf8..e281c96474a1 100644 --- a/src/kernel/type_checker.h +++ b/src/kernel/type_checker.h @@ -42,6 +42,7 @@ class type_checker { private: bool m_st_owner; state * m_st; + diagnostics * m_diag; local_ctx m_lctx; definition_safety m_definition_safety; /* When `m_lparams != nullptr, the `check` method makes sure all level parameters @@ -104,8 +105,8 @@ class type_checker { public: type_checker(state & st, local_ctx const & lctx, definition_safety ds = definition_safety::safe); type_checker(state & st, definition_safety ds = definition_safety::safe):type_checker(st, local_ctx(), ds) {} - type_checker(environment const & env, local_ctx const & lctx, definition_safety ds = definition_safety::safe); - type_checker(environment const & env, definition_safety ds = definition_safety::safe):type_checker(env, local_ctx(), ds) {} + type_checker(environment const & env, local_ctx const & lctx, diagnostics * diag = nullptr, definition_safety ds = definition_safety::safe); + type_checker(environment const & env, diagnostics * diag = nullptr, definition_safety ds = definition_safety::safe):type_checker(env, local_ctx(), diag, ds) {} type_checker(type_checker &&); type_checker(type_checker const &) = delete; ~type_checker(); diff --git a/tests/lean/run/ack.lean b/tests/lean/run/ack.lean index e86cc92c8a7e..41fd2ed6ea9a 100644 --- a/tests/lean/run/ack.lean +++ b/tests/lean/run/ack.lean @@ -3,3 +3,32 @@ def ack : Nat → Nat → Nat | x+1, 0 => ack x 1 | x+1, y+1 => ack x (ack (x+1) y) termination_by a b => (a, b) + +/-- +info: [reduction] unfolded declarations (max: 1725, num: 4): + Nat.rec ↦ 1725 + ⏎ + Eq.rec ↦ 1114 + ⏎ + Acc.rec ↦ 1050 + ⏎ + PSigma.rec ↦ 513[reduction] unfolded reducible declarations (max: 1577, num: 3): + Nat.casesOn ↦ 1577 + ⏎ + Eq.ndrec ↦ 984 + ⏎ + PSigma.casesOn ↦ 513[kernel] unfolded declarations (max: 1193, num: 5): + Nat.casesOn ↦ 1193 + ⏎ + Nat.rec ↦ 1065 + ⏎ + Eq.ndrec ↦ 973 + ⏎ + Eq.rec ↦ 973 + Acc.rec ↦ 754use `set_option diagnostics.threshold ` to control threshold for reporting counters +-/ +#guard_msgs in +set_option diagnostics.threshold 500 in +set_option diagnostics true in +theorem ex : ack 3 2 = 29 := + rfl From dfde4ee3aa2f38be1769d5719e52534cc7d4f5bf Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Mon, 6 May 2024 15:50:18 -0700 Subject: [PATCH 06/44] chore: update stage0 --- stage0/src/CMakeLists.txt | 10 +- stage0/src/kernel/declaration.cpp | 8 +- stage0/src/kernel/declaration.h | 7 +- stage0/src/kernel/environment.cpp | 82 +- stage0/src/kernel/environment.h | 33 + stage0/src/kernel/inductive.cpp | 14 +- stage0/src/kernel/type_checker.cpp | 19 +- stage0/src/kernel/type_checker.h | 5 +- .../src/library/compiler/ir_interpreter.cpp | 2 + stage0/src/runtime/process.cpp | 36 + stage0/stdlib/Init/Control/Lawful/Basic.c | 410 +- stage0/stdlib/Init/Core.c | 27 - stage0/stdlib/Init/Data/Array/Basic.c | 6 +- stage0/stdlib/Init/Data/Array/Lemmas.c | 365 + stage0/stdlib/Init/Data/Array/Mem.c | 803 +- stage0/stdlib/Init/Data/List.c | 6 +- stage0/stdlib/Init/Data/List/BasicAux.c | 430 +- stage0/stdlib/Init/Data/List/Impl.c | 76 - stage0/stdlib/Init/Data/List/Lemmas.c | 244 +- stage0/stdlib/Init/Data/List/TakeDrop.c | 97 + stage0/stdlib/Init/Data/String/Basic.c | 12 +- stage0/stdlib/Init/Data/String/Extra.c | 249 +- stage0/stdlib/Init/Meta.c | 1876 +- stage0/stdlib/Init/Notation.c | 648 + stage0/stdlib/Init/Prelude.c | 2 +- stage0/stdlib/Init/System/IO.c | 21 + stage0/stdlib/Init/Tactics.c | 65 +- stage0/stdlib/Init/WFTactics.c | 378 +- stage0/stdlib/Lean/BuiltinDocAttr.c | 4 +- stage0/stdlib/Lean/Compiler/ExternAttr.c | 738 +- stage0/stdlib/Lean/Compiler/IR/Basic.c | 2 +- stage0/stdlib/Lean/Compiler/IR/RC.c | 6 +- stage0/stdlib/Lean/Compiler/IR/ResetReuse.c | 2695 ++- stage0/stdlib/Lean/Compiler/InitAttr.c | 90 +- .../stdlib/Lean/Compiler/LCNF/AuxDeclCache.c | 14 +- stage0/stdlib/Lean/Compiler/LCNF/Basic.c | 38 +- stage0/stdlib/Lean/Compiler/LCNF/Closure.c | 3100 +-- stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c | 2 +- stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c | 2 +- stage0/stdlib/Lean/Compiler/LCNF/Main.c | 2218 +- .../stdlib/Lean/Compiler/LCNF/PassManager.c | 113 +- stage0/stdlib/Lean/Compiler/LCNF/Passes.c | 34 +- stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c | 66 +- .../stdlib/Lean/Compiler/LCNF/PrettyPrinter.c | 314 +- .../stdlib/Lean/Compiler/LCNF/Simp/JpCases.c | 202 +- stage0/stdlib/Lean/Compiler/LCNF/Simp/Main.c | 10 +- stage0/stdlib/Lean/Compiler/LCNF/Simp/SimpM.c | 10 +- stage0/stdlib/Lean/Compiler/LCNF/Specialize.c | 158 +- stage0/stdlib/Lean/Compiler/LCNF/Testing.c | 4 +- stage0/stdlib/Lean/Compiler/Main.c | 4 +- stage0/stdlib/Lean/CoreM.c | 10739 ++++++---- stage0/stdlib/Lean/Data/HashMap.c | 2 +- stage0/stdlib/Lean/Data/Lsp/Capabilities.c | 16 +- .../stdlib/Lean/Data/Lsp/LanguageFeatures.c | 9234 ++++---- stage0/stdlib/Lean/Declaration.c | 82 +- stage0/stdlib/Lean/Elab/App.c | 158 +- stage0/stdlib/Lean/Elab/Binders.c | 3035 ++- stage0/stdlib/Lean/Elab/BuiltinNotation.c | 1206 +- stage0/stdlib/Lean/Elab/BuiltinTerm.c | 974 +- stage0/stdlib/Lean/Elab/Command.c | 4922 +++-- stage0/stdlib/Lean/Elab/DefView.c | 8 +- stage0/stdlib/Lean/Elab/Deriving/BEq.c | 43 +- stage0/stdlib/Lean/Elab/Deriving/DecEq.c | 47 +- stage0/stdlib/Lean/Elab/Deriving/FromToJson.c | 353 +- stage0/stdlib/Lean/Elab/Deriving/Hashable.c | 43 +- stage0/stdlib/Lean/Elab/Deriving/Inhabited.c | 917 +- stage0/stdlib/Lean/Elab/Deriving/Ord.c | 43 +- stage0/stdlib/Lean/Elab/Deriving/Repr.c | 43 +- stage0/stdlib/Lean/Elab/GuardMsgs.c | 4 +- stage0/stdlib/Lean/Elab/Inductive.c | 1064 +- stage0/stdlib/Lean/Elab/InfoTree/Main.c | 1158 +- stage0/stdlib/Lean/Elab/Match.c | 1245 +- stage0/stdlib/Lean/Elab/MutualDef.c | 718 +- stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c | 3021 +-- .../Elab/PreDefinition/Structural/BRecOn.c | 8 +- .../Lean/Elab/PreDefinition/Structural/Eqns.c | 932 +- .../Elab/PreDefinition/Structural/IndPred.c | 4 +- .../PreDefinition/Structural/SmartUnfolding.c | 217 +- .../stdlib/Lean/Elab/PreDefinition/WF/Eqns.c | 1523 +- .../Lean/Elab/PreDefinition/WF/GuessLex.c | 10 +- .../Lean/Elab/PreDefinition/WF/PackMutual.c | 982 +- stage0/stdlib/Lean/Elab/Structure.c | 794 +- stage0/stdlib/Lean/Elab/Syntax.c | 4 +- stage0/stdlib/Lean/Elab/Tactic/Basic.c | 1050 +- .../stdlib/Lean/Elab/Tactic/BuiltinTactic.c | 1067 +- stage0/stdlib/Lean/Elab/Tactic/Conv/Pattern.c | 644 +- stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c | 50 +- stage0/stdlib/Lean/Elab/Tactic/Ext.c | 6 +- stage0/stdlib/Lean/Elab/Tactic/NormCast.c | 4608 ++-- stage0/stdlib/Lean/Elab/Tactic/RCases.c | 4 +- stage0/stdlib/Lean/Elab/Tactic/Simp.c | 1922 +- stage0/stdlib/Lean/Elab/Tactic/SimpTrace.c | 2620 +-- stage0/stdlib/Lean/Elab/Tactic/Simpa.c | 2231 +- stage0/stdlib/Lean/Elab/Tactic/Simproc.c | 4 + stage0/stdlib/Lean/Elab/Term.c | 3979 +++- stage0/stdlib/Lean/Environment.c | 1071 +- stage0/stdlib/Lean/KeyedDeclsAttribute.c | 2 - stage0/stdlib/Lean/Linter/MissingDocs.c | 10 +- stage0/stdlib/Lean/Linter/UnusedVariables.c | 12 +- stage0/stdlib/Lean/LocalContext.c | 12 +- stage0/stdlib/Lean/Meta.c | 6 +- stage0/stdlib/Lean/Meta/Basic.c | 10345 ++++----- stage0/stdlib/Lean/Meta/Check.c | 804 +- stage0/stdlib/Lean/Meta/CoeAttr.c | 688 +- stage0/stdlib/Lean/Meta/Constructions.c | 6 +- stage0/stdlib/Lean/Meta/Diagnostics.c | 3565 ++++ stage0/stdlib/Lean/Meta/Eqns.c | 76 +- stage0/stdlib/Lean/Meta/ExprDefEq.c | 1972 +- stage0/stdlib/Lean/Meta/IndPredBelow.c | 83 +- stage0/stdlib/Lean/Meta/Injective.c | 4462 +++- stage0/stdlib/Lean/Meta/Instances.c | 1569 +- stage0/stdlib/Lean/Meta/LazyDiscrTree.c | 1838 +- stage0/stdlib/Lean/Meta/Match/Match.c | 36 +- stage0/stdlib/Lean/Meta/Match/MatchEqs.c | 6 +- .../Lean/Meta/Match/MatcherApp/Transform.c | 54 +- stage0/stdlib/Lean/Meta/SynthInstance.c | 3552 +++- stage0/stdlib/Lean/Meta/Tactic/AC/Main.c | 252 +- stage0/stdlib/Lean/Meta/Tactic/Acyclic.c | 517 +- stage0/stdlib/Lean/Meta/Tactic/Backtrack.c | 1483 +- stage0/stdlib/Lean/Meta/Tactic/FunInd.c | 388 +- .../stdlib/Lean/Meta/Tactic/LibrarySearch.c | 874 +- .../Lean/Meta/Tactic/LinearArith/Solver.c | 602 +- stage0/stdlib/Lean/Meta/Tactic/NormCast.c | 1258 +- stage0/stdlib/Lean/Meta/Tactic/Rfl.c | 60 +- stage0/stdlib/Lean/Meta/Tactic/Simp.c | 6 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c | 64 +- .../Meta/Tactic/Simp/BuiltinSimprocs/Nat.c | 4 +- .../Lean/Meta/Tactic/Simp/Diagnostics.c | 1955 ++ stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c | 10408 ++++----- stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c | 6154 +++--- stage0/stdlib/Lean/Meta/Tactic/Simp/SimpAll.c | 3015 +-- .../Lean/Meta/Tactic/Simp/SimpTheorems.c | 361 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c | 34 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Types.c | 4103 ++-- stage0/stdlib/Lean/Meta/Tactic/SolveByElim.c | 10 +- stage0/stdlib/Lean/Meta/Tactic/Split.c | 334 +- stage0/stdlib/Lean/Meta/Tactic/SplitIf.c | 217 +- stage0/stdlib/Lean/Meta/Tactic/TryThis.c | 14 +- stage0/stdlib/Lean/Meta/Tactic/Unfold.c | 240 +- stage0/stdlib/Lean/Meta/Tactic/Util.c | 1243 +- stage0/stdlib/Lean/Meta/WHNF.c | 1600 +- stage0/stdlib/Lean/Parser/Extension.c | 41 +- stage0/stdlib/Lean/Parser/Term.c | 106 +- stage0/stdlib/Lean/PrettyPrinter.c | 2874 ++- .../Lean/PrettyPrinter/Delaborator/Basic.c | 424 +- .../Lean/PrettyPrinter/Delaborator/Builtins.c | 17406 ++++++++-------- .../PrettyPrinter/Delaborator/FieldNotation.c | 345 +- .../Delaborator/TopDownAnalyze.c | 800 +- stage0/stdlib/Lean/ReducibilityAttrs.c | 3324 ++- stage0/stdlib/Lean/Replay.c | 608 +- stage0/stdlib/Lean/Server/CodeActions/Attr.c | 17 + stage0/stdlib/Lean/Server/CodeActions/Basic.c | 4 + stage0/stdlib/Lean/Server/Completion.c | 54 +- .../stdlib/Lean/Server/CompletionItemData.c | 8 +- stage0/stdlib/Lean/Server/FileWorker.c | 10 +- .../Lean/Server/FileWorker/RequestHandling.c | 178 +- .../Lean/Server/FileWorker/WidgetRequests.c | 2 - stage0/stdlib/Lean/Server/ImportCompletion.c | 4 +- stage0/stdlib/Lean/Server/InfoUtils.c | 12 - stage0/stdlib/Lean/Server/References.c | 6 +- stage0/stdlib/Lean/Server/Watchdog.c | 100 +- stage0/stdlib/Lean/SubExpr.c | 474 +- stage0/stdlib/Lean/ToExpr.c | 273 + stage0/stdlib/Lean/Util/ForEachExprWhere.c | 672 +- stage0/stdlib/Lean/Util/Trace.c | 2462 ++- .../Lean/Widget/InteractiveDiagnostic.c | 6 +- stage0/stdlib/Lean/Widget/UserWidget.c | 16 +- 167 files changed, 104147 insertions(+), 68933 deletions(-) create mode 100644 stage0/stdlib/Init/Data/List/TakeDrop.c create mode 100644 stage0/stdlib/Lean/Meta/Diagnostics.c create mode 100644 stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c diff --git a/stage0/src/CMakeLists.txt b/stage0/src/CMakeLists.txt index fe86b332d86c..ee08a831f339 100644 --- a/stage0/src/CMakeLists.txt +++ b/stage0/src/CMakeLists.txt @@ -9,7 +9,7 @@ endif() include(ExternalProject) project(LEAN CXX C) set(LEAN_VERSION_MAJOR 4) -set(LEAN_VERSION_MINOR 8) +set(LEAN_VERSION_MINOR 9) set(LEAN_VERSION_PATCH 0) set(LEAN_VERSION_IS_RELEASE 0) # This number is 1 in the release revision, and 0 otherwise. set(LEAN_SPECIAL_VERSION_DESC "" CACHE STRING "Additional version description like 'nightly-2018-03-11'") @@ -315,6 +315,12 @@ endif() string(APPEND TOOLCHAIN_STATIC_LINKER_FLAGS " ${LEAN_CXX_STDLIB}") string(APPEND TOOLCHAIN_SHARED_LINKER_FLAGS " ${LEAN_CXX_STDLIB}") +# in local builds, link executables and not just dynlibs against C++ stdlib as well, +# which is required for e.g. asan +if(NOT LEAN_STANDALONE) + string(APPEND CMAKE_EXE_LINKER_FLAGS " ${LEAN_CXX_STDLIB}") +endif() + # flags for user binaries = flags for toolchain binaries + Lake string(APPEND LEANC_STATIC_LINKER_FLAGS " ${TOOLCHAIN_STATIC_LINKER_FLAGS} -lLake") @@ -585,7 +591,7 @@ endif() if(PREV_STAGE) add_custom_target(update-stage0 - COMMAND bash -c 'CSRCS=${CMAKE_BINARY_DIR}/lib/temp script/update-stage0' + COMMAND bash -c 'CSRCS=${CMAKE_BINARY_DIR}/lib/temp script/lib/update-stage0' DEPENDS make_stdlib WORKING_DIRECTORY "${LEAN_SOURCE_DIR}/..") diff --git a/stage0/src/kernel/declaration.cpp b/stage0/src/kernel/declaration.cpp index cf7f84aa2140..422b1180b342 100644 --- a/stage0/src/kernel/declaration.cpp +++ b/stage0/src/kernel/declaration.cpp @@ -71,8 +71,10 @@ definition_val::definition_val(name const & n, names const & lparams, expr const definition_safety definition_val::get_safety() const { return static_cast(lean_definition_val_get_safety(to_obj_arg())); } -theorem_val::theorem_val(name const & n, names const & lparams, expr const & type, expr const & val): - object_ref(mk_cnstr(0, constant_val(n, lparams, type), val)) { +extern "C" object * lean_mk_theorem_val(object * n, object * lparams, object * type, object * value, object * all); + +theorem_val::theorem_val(name const & n, names const & lparams, expr const & type, expr const & val, names const & all): + object_ref(lean_mk_theorem_val(n.to_obj_arg(), lparams.to_obj_arg(), type.to_obj_arg(), val.to_obj_arg(), all.to_obj_arg())) { } extern "C" object * lean_mk_opaque_val(object * n, object * lparams, object * type, object * value, uint8 is_unsafe, object * all); @@ -206,7 +208,7 @@ declaration mk_definition(environment const & env, name const & n, names const & } declaration mk_theorem(name const & n, names const & lparams, expr const & type, expr const & val) { - return declaration(mk_cnstr(static_cast(declaration_kind::Theorem), theorem_val(n, lparams, type, val))); + return declaration(mk_cnstr(static_cast(declaration_kind::Theorem), theorem_val(n, lparams, type, val, names(n)))); } declaration mk_opaque(name const & n, names const & params, expr const & t, expr const & v, bool is_unsafe) { diff --git a/stage0/src/kernel/declaration.h b/stage0/src/kernel/declaration.h index 7d67aaa0d2c3..e8513efcd289 100644 --- a/stage0/src/kernel/declaration.h +++ b/stage0/src/kernel/declaration.h @@ -112,12 +112,13 @@ class definition_val : public object_ref { typedef list_ref definition_vals; /* -structure theorem_val extends constant_val := -(value : task expr) +structure TheoremVal extends ConstantVal where + value : Expr + all : List Name := [name] */ class theorem_val : public object_ref { public: - theorem_val(name const & n, names const & lparams, expr const & type, expr const & val); + theorem_val(name const & n, names const & lparams, expr const & type, expr const & val, names const & all); theorem_val(theorem_val const & other):object_ref(other) {} theorem_val(theorem_val && other):object_ref(other) {} theorem_val & operator=(theorem_val const & other) { object_ref::operator=(other); return *this; } diff --git a/stage0/src/kernel/environment.cpp b/stage0/src/kernel/environment.cpp index 792f4880dfe3..d8316d5c9f3e 100644 --- a/stage0/src/kernel/environment.cpp +++ b/stage0/src/kernel/environment.cpp @@ -28,6 +28,38 @@ extern "C" object* lean_get_extension(object*, object*); extern "C" object* lean_set_extension(object*, object*, object*); extern "C" object* lean_environment_set_main_module(object*, object*); extern "C" object* lean_environment_main_module(object*); +extern "C" object* lean_kernel_record_unfold (object*, object*); +extern "C" object* lean_kernel_get_diag(object*); +extern "C" object* lean_kernel_set_diag(object*, object*); +extern "C" uint8* lean_kernel_diag_is_enabled(object*); + +void diagnostics::record_unfold(name const & decl_name) { + m_obj = lean_kernel_record_unfold(to_obj_arg(), decl_name.to_obj_arg()); +} + +scoped_diagnostics::scoped_diagnostics(environment const & env, bool collect) { + if (collect) { + diagnostics d(env.get_diag()); + if (lean_kernel_diag_is_enabled(d.to_obj_arg())) { + m_diag = new diagnostics(d); + } else + m_diag = nullptr; + } else { + m_diag = nullptr; + } +} + +scoped_diagnostics::~scoped_diagnostics() { + if (m_diag) + delete m_diag; +} + +environment scoped_diagnostics::update(environment const & env) const { + if (m_diag) + return env.set_diag(*m_diag); + else + return env; +} environment mk_empty_environment(uint32 trust_lvl) { return get_io_result(lean_mk_empty_environment(trust_lvl, io_mk_world())); @@ -37,6 +69,14 @@ environment::environment(unsigned trust_lvl): object_ref(mk_empty_environment(trust_lvl)) { } +diagnostics environment::get_diag() const { + return diagnostics(lean_kernel_get_diag(to_obj_arg())); +} + +environment environment::set_diag(diagnostics const & diag) const { + return environment(lean_kernel_set_diag(to_obj_arg(), diag.to_obj_arg())); +} + void environment::set_main_module(name const & n) { m_obj = lean_environment_set_main_module(m_obj, n.to_obj_arg()); } @@ -118,13 +158,13 @@ static void check_constant_val(environment const & env, constant_val const & v, checker.ensure_sort(sort, v.get_type()); } -static void check_constant_val(environment const & env, constant_val const & v, definition_safety ds) { - type_checker checker(env, ds); +static void check_constant_val(environment const & env, constant_val const & v, diagnostics * diag, definition_safety ds) { + type_checker checker(env, diag, ds); check_constant_val(env, v, checker); } -static void check_constant_val(environment const & env, constant_val const & v, bool safe_only) { - check_constant_val(env, v, safe_only ? definition_safety::safe : definition_safety::unsafe); +static void check_constant_val(environment const & env, constant_val const & v, diagnostics * diag, bool safe_only) { + check_constant_val(env, v, diag, safe_only ? definition_safety::safe : definition_safety::unsafe); } void environment::add_core(constant_info const & info) { @@ -136,48 +176,50 @@ environment environment::add(constant_info const & info) const { } environment environment::add_axiom(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); axiom_val const & v = d.to_axiom_val(); if (check) - check_constant_val(*this, v.to_constant_val(), !d.is_unsafe()); - return add(constant_info(d)); + check_constant_val(*this, v.to_constant_val(), diag.get(), !d.is_unsafe()); + return diag.update(add(constant_info(d))); } environment environment::add_definition(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); definition_val const & v = d.to_definition_val(); if (v.is_unsafe()) { /* Meta definition can be recursive. So, we check the header, add, and then type check the body. */ if (check) { - type_checker checker(*this, definition_safety::unsafe); + type_checker checker(*this, diag.get(), definition_safety::unsafe); check_constant_val(*this, v.to_constant_val(), checker); } environment new_env = add(constant_info(d)); if (check) { - type_checker checker(new_env, definition_safety::unsafe); + type_checker checker(new_env, diag.get(), definition_safety::unsafe); check_no_metavar_no_fvar(new_env, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(new_env, d, val_type); } - return new_env; + return diag.update(new_env); } else { if (check) { - type_checker checker(*this); + type_checker checker(*this, diag.get()); 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()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } } environment environment::add_theorem(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); theorem_val const & v = d.to_theorem_val(); if (check) { - // TODO(Leo): we must add support for handling tasks here - type_checker checker(*this); + type_checker checker(*this, diag.get()); 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); @@ -186,22 +228,24 @@ environment environment::add_theorem(declaration const & d, bool check) const { if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } environment environment::add_opaque(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); opaque_val const & v = d.to_opaque_val(); if (check) { - type_checker checker(*this); + type_checker checker(*this, diag.get()); check_constant_val(*this, v.to_constant_val(), checker); expr val_type = checker.check(v.get_value(), v.get_lparams()); if (!checker.is_def_eq(val_type, v.get_type())) throw definition_type_mismatch_exception(*this, d, val_type); } - return add(constant_info(d)); + return diag.update(add(constant_info(d))); } environment environment::add_mutual(declaration const & d, bool check) const { + scoped_diagnostics diag(*this, check); definition_vals const & vs = d.to_definition_vals(); if (empty(vs)) throw kernel_exception(*this, "invalid empty mutual definition"); @@ -210,7 +254,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { throw kernel_exception(*this, "invalid mutual definition, declaration is not tagged as unsafe/partial"); /* Check declarations header */ if (check) { - type_checker checker(*this, safety); + type_checker checker(*this, diag.get(), safety); for (definition_val const & v : vs) { if (v.get_safety() != safety) throw kernel_exception(*this, "invalid mutual definition, declarations must have the same safety annotation"); @@ -224,7 +268,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { } /* Check actual definitions */ if (check) { - type_checker checker(new_env, safety); + type_checker checker(new_env, diag.get(), safety); for (definition_val const & v : vs) { check_no_metavar_no_fvar(new_env, v.get_name(), v.get_value()); expr val_type = checker.check(v.get_value(), v.get_lparams()); @@ -232,7 +276,7 @@ environment environment::add_mutual(declaration const & d, bool check) const { throw definition_type_mismatch_exception(new_env, d, val_type); } } - return new_env; + return diag.update(new_env); } environment environment::add(declaration const & d, bool check) const { diff --git a/stage0/src/kernel/environment.h b/stage0/src/kernel/environment.h index d9cb69f47258..eaa077dbf16a 100644 --- a/stage0/src/kernel/environment.h +++ b/stage0/src/kernel/environment.h @@ -29,6 +29,36 @@ class environment_extension { virtual ~environment_extension() {} }; +/* Wrapper for `Kernel.Diagnostics` */ +class diagnostics : public object_ref { +public: + diagnostics(diagnostics const & other):object_ref(other) {} + diagnostics(diagnostics && other):object_ref(other) {} + explicit diagnostics(b_obj_arg o, bool b):object_ref(o, b) {} + explicit diagnostics(obj_arg o):object_ref(o) {} + ~diagnostics() {} + void record_unfold(name const & decl_name); +}; + +/* +Store `Kernel.Diagnostics` stored in environment extension in `m_diag` IF +- `Kernel.Diagnostics.enable = true` +- `collect = true`. This is a minor optimization. + +We use this class to ensure we don't waste time collecting information +that was not requested. +*/ +class scoped_diagnostics { + diagnostics * m_diag; +public: + scoped_diagnostics(environment const & env, bool collect); + scoped_diagnostics(scoped_diagnostics const &) = delete; + scoped_diagnostics(scoped_diagnostics &&) = delete; + ~scoped_diagnostics(); + environment update(environment const &) const; + diagnostics * get() const { return m_diag; } +}; + class environment : public object_ref { friend class add_inductive_fn; @@ -56,6 +86,9 @@ class environment : public object_ref { environment & operator=(environment const & other) { object_ref::operator=(other); return *this; } environment & operator=(environment && other) { object_ref::operator=(other); return *this; } + diagnostics get_diag() const; + environment set_diag(diagnostics const & diag) const; + /** \brief Return the trust level of this environment. */ unsigned trust_lvl() const; diff --git a/stage0/src/kernel/inductive.cpp b/stage0/src/kernel/inductive.cpp index 2238b24a256c..8b75cdeff2a7 100644 --- a/stage0/src/kernel/inductive.cpp +++ b/stage0/src/kernel/inductive.cpp @@ -124,6 +124,7 @@ optional get_rec_rule_for(recursor_val const & rec_val, expr cons class add_inductive_fn { environment m_env; name_generator m_ngen; + diagnostics * m_diag; local_ctx m_lctx; names m_lparams; unsigned m_nparams; @@ -158,8 +159,8 @@ class add_inductive_fn { buffer m_rec_infos; public: - add_inductive_fn(environment const & env, inductive_decl const & decl, bool is_nested): - m_env(env), m_ngen(*g_ind_fresh), m_lparams(decl.get_lparams()), m_is_unsafe(decl.is_unsafe()), + add_inductive_fn(environment const & env, diagnostics * diag, inductive_decl const & decl, bool is_nested): + m_env(env), m_ngen(*g_ind_fresh), m_diag(diag), m_lparams(decl.get_lparams()), m_is_unsafe(decl.is_unsafe()), m_is_nested(is_nested) { if (!decl.get_nparams().is_small()) throw kernel_exception(env, "invalid inductive datatype, number of parameters is too big"); @@ -167,7 +168,7 @@ class add_inductive_fn { to_buffer(decl.get_types(), m_ind_types); } - type_checker tc() { return type_checker(m_env, m_lctx, m_is_unsafe ? definition_safety::unsafe : definition_safety::safe); } + type_checker tc() { return type_checker(m_env, m_lctx, m_diag, m_is_unsafe ? definition_safety::unsafe : definition_safety::safe); } /** Return type of the parameter at position `i` */ expr get_param_type(unsigned i) const { @@ -1112,10 +1113,11 @@ static pair> mk_aux_rec_name_map(environment const & aux_e environment environment::add_inductive(declaration const & d) const { elim_nested_inductive_result res = elim_nested_inductive_fn(*this, d)(); bool is_nested = !res.m_aux2nested.empty(); - environment aux_env = add_inductive_fn(*this, inductive_decl(res.m_aux_decl), is_nested)(); + scoped_diagnostics diag(*this, true); + environment aux_env = add_inductive_fn(*this, diag.get(), inductive_decl(res.m_aux_decl), is_nested)(); if (!is_nested) { /* `d` did not contain nested inductive types. */ - return aux_env; + return diag.update(aux_env); } else { /* Restore nested inductives. */ inductive_decl ind_d(d); @@ -1170,7 +1172,7 @@ environment environment::add_inductive(declaration const & d) const { for (name const & aux_rec : aux_rec_names) { process_rec(aux_rec); } - return new_env; + return diag.update(new_env); } } diff --git a/stage0/src/kernel/type_checker.cpp b/stage0/src/kernel/type_checker.cpp index 9adb025818c2..1dbb0819a98e 100644 --- a/stage0/src/kernel/type_checker.cpp +++ b/stage0/src/kernel/type_checker.cpp @@ -477,6 +477,11 @@ expr type_checker::whnf_core(expr const & e, bool cheap_rec, bool cheap_proj) { cheap_rec, cheap_proj); } else if (f == f0) { if (auto r = reduce_recursor(e, cheap_rec, cheap_proj)) { + if (m_diag) { + auto f = get_app_fn(e); + if (is_constant(f)) + m_diag->record_unfold(const_name(f)); + } /* iota-reduction and quotient reduction rules */ return whnf_core(*r, cheap_rec, cheap_proj); } else { @@ -513,8 +518,12 @@ optional type_checker::is_delta(expr const & e) const { optional type_checker::unfold_definition_core(expr const & e) { if (is_constant(e)) { if (auto d = is_delta(e)) { - if (length(const_levels(e)) == d->get_num_lparams()) + if (length(const_levels(e)) == d->get_num_lparams()) { + if (m_diag) { + m_diag->record_unfold(d->get_name()); + } return some_expr(instantiate_value_lparams(*d, const_levels(e))); + } } } return none_expr(); @@ -1148,18 +1157,18 @@ expr type_checker::eta_expand(expr const & e) { return m_lctx.mk_lambda(fvars, r); } -type_checker::type_checker(environment const & env, local_ctx const & lctx, definition_safety ds): - m_st_owner(true), m_st(new state(env)), +type_checker::type_checker(environment const & env, local_ctx const & lctx, diagnostics * diag, definition_safety ds): + m_st_owner(true), m_st(new state(env)), m_diag(diag), m_lctx(lctx), m_definition_safety(ds), m_lparams(nullptr) { } type_checker::type_checker(state & st, local_ctx const & lctx, definition_safety ds): - m_st_owner(false), m_st(&st), m_lctx(lctx), + m_st_owner(false), m_st(&st), m_diag(nullptr), m_lctx(lctx), m_definition_safety(ds), m_lparams(nullptr) { } type_checker::type_checker(type_checker && src): - m_st_owner(src.m_st_owner), m_st(src.m_st), m_lctx(std::move(src.m_lctx)), + m_st_owner(src.m_st_owner), m_st(src.m_st), m_diag(src.m_diag), m_lctx(std::move(src.m_lctx)), m_definition_safety(src.m_definition_safety), m_lparams(src.m_lparams) { src.m_st_owner = false; } diff --git a/stage0/src/kernel/type_checker.h b/stage0/src/kernel/type_checker.h index e38a60772bf8..e281c96474a1 100644 --- a/stage0/src/kernel/type_checker.h +++ b/stage0/src/kernel/type_checker.h @@ -42,6 +42,7 @@ class type_checker { private: bool m_st_owner; state * m_st; + diagnostics * m_diag; local_ctx m_lctx; definition_safety m_definition_safety; /* When `m_lparams != nullptr, the `check` method makes sure all level parameters @@ -104,8 +105,8 @@ class type_checker { public: type_checker(state & st, local_ctx const & lctx, definition_safety ds = definition_safety::safe); type_checker(state & st, definition_safety ds = definition_safety::safe):type_checker(st, local_ctx(), ds) {} - type_checker(environment const & env, local_ctx const & lctx, definition_safety ds = definition_safety::safe); - type_checker(environment const & env, definition_safety ds = definition_safety::safe):type_checker(env, local_ctx(), ds) {} + type_checker(environment const & env, local_ctx const & lctx, diagnostics * diag = nullptr, definition_safety ds = definition_safety::safe); + type_checker(environment const & env, diagnostics * diag = nullptr, definition_safety ds = definition_safety::safe):type_checker(env, local_ctx(), diag, ds) {} type_checker(type_checker &&); type_checker(type_checker const &) = delete; ~type_checker(); diff --git a/stage0/src/library/compiler/ir_interpreter.cpp b/stage0/src/library/compiler/ir_interpreter.cpp index ce6698ac11fb..cd88e35ba6af 100644 --- a/stage0/src/library/compiler/ir_interpreter.cpp +++ b/stage0/src/library/compiler/ir_interpreter.cpp @@ -948,6 +948,8 @@ class interpreter { m_prefer_native = opts.get_bool(*g_interpreter_prefer_native, LEAN_DEFAULT_INTERPRETER_PREFER_NATIVE); } + interpreter(interpreter const &) = delete; + ~interpreter() { for_each(m_constant_cache, [](name const &, constant_cache_entry const & e) { if (!e.m_is_scalar) { diff --git a/stage0/src/runtime/process.cpp b/stage0/src/runtime/process.cpp index 5f45e3feb1c0..035442d1868d 100644 --- a/stage0/src/runtime/process.cpp +++ b/stage0/src/runtime/process.cpp @@ -24,6 +24,7 @@ Author: Jared Roesch #include #include #include +#include // NOLINT #endif #include "runtime/object.h" @@ -57,6 +58,24 @@ lean_object * wrap_win_handle(HANDLE h) { return lean_alloc_external(g_win_handle_external_class, static_cast(h)); } +extern "C" LEAN_EXPORT obj_res lean_io_process_get_current_dir(obj_arg) { + char path[MAX_PATH]; + DWORD sz = GetCurrentDirectory(MAX_PATH, path); + if (sz != 0) { + return io_result_mk_ok(lean_mk_string_from_bytes(path, sz)); + } else { + return io_result_mk_error((sstream() << GetLastError()).str()); + } +} + +extern "C" LEAN_EXPORT obj_res lean_io_process_set_current_dir(b_obj_arg path, obj_arg) { + if (SetCurrentDirectory(string_cstr(path))) { + return io_result_mk_ok(box(0)); + } else { + return io_result_mk_error((sstream() << GetLastError()).str()); + } +} + extern "C" LEAN_EXPORT obj_res lean_io_process_get_pid(obj_arg) { return lean_io_result_mk_ok(box_uint32(GetCurrentProcessId())); } @@ -252,6 +271,23 @@ void finalize_process() {} #else +extern "C" LEAN_EXPORT obj_res lean_io_process_get_current_dir(obj_arg) { + char path[PATH_MAX]; + if (getcwd(path, PATH_MAX)) { + return io_result_mk_ok(mk_string(path)); + } else { + return io_result_mk_error(decode_io_error(errno, nullptr)); + } +} + +extern "C" LEAN_EXPORT obj_res lean_io_process_set_current_dir(b_obj_arg path, obj_arg) { + if (!chdir(string_cstr(path))) { + return io_result_mk_ok(box(0)); + } else { + return io_result_mk_error(decode_io_error(errno, path)); + } +} + extern "C" LEAN_EXPORT obj_res lean_io_process_get_pid(obj_arg) { static_assert(sizeof(pid_t) == sizeof(uint32), "pid_t is expected to be a 32-bit type"); // NOLINT return lean_io_result_mk_ok(box_uint32(getpid())); diff --git a/stage0/stdlib/Init/Control/Lawful/Basic.c b/stage0/stdlib/Init/Control/Lawful/Basic.c index 8aba895d0192..9a28244674ab 100644 --- a/stage0/stdlib/Init/Control/Lawful/Basic.c +++ b/stage0/stdlib/Init/Control/Lawful/Basic.c @@ -13,49 +13,49 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1751_; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__26; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__8; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1821_; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__23; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__12; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__5; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__26; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__30; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__27; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1673_; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613_; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__23; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1891_; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11; -static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18; -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1() { +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24; +static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1() { _start: { lean_object* x_1; @@ -63,7 +63,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2() { _start: { lean_object* x_1; @@ -71,7 +71,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3() { _start: { lean_object* x_1; @@ -79,7 +79,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4() { _start: { lean_object* x_1; @@ -87,19 +87,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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_Init_Control_Lawful_Basic___hyg_1613____closed__1; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3; -x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; +x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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_Init_Control_Lawful_Basic___hyg_1613____closed__6() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -108,7 +108,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7() { _start: { lean_object* x_1; @@ -116,19 +116,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3; -x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; +x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9() { _start: { lean_object* x_1; @@ -136,17 +136,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11() { _start: { lean_object* x_1; @@ -154,47 +154,47 @@ x_1 = lean_mk_string_from_bytes("intros", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3; -x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; +x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -202,23 +202,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__12; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -226,17 +226,17 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19() { _start: { lean_object* x_1; @@ -244,29 +244,29 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22() { _start: { lean_object* x_1; @@ -274,19 +274,19 @@ x_1 = lean_mk_string_from_bytes("tacticRfl", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__23() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3; -x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; +x_4 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24() { _start: { lean_object* x_1; @@ -294,35 +294,35 @@ x_1 = lean_mk_string_from_bytes("rfl", 3); return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24; 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_Init_Control_Lawful_Basic___hyg_1613____closed__26() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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_Init_Control_Lawful_Basic___hyg_1613____closed__23; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__26; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__23; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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); @@ -330,23 +330,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__27; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29() { _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_Init_Control_Lawful_Basic___hyg_1613____closed__10; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -354,23 +354,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__8; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____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); @@ -378,23 +378,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6; -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = lean_box(2); -x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5; -x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32; +x_2 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__5; +x_3 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32; x_4 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -402,43 +402,43 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1673_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1751_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1821_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1891_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33; +x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } @@ -455,82 +455,82 @@ lean_dec_ref(res); res = initialize_Init_Meta(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__1); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__2); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__3); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__4); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__5); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__6); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__7); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__8); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__9); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__10); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__11); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__12); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__13); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__14); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__15); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__16); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__17); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__18); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__19); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__20); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__21); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__22); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__23 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__23(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__23); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__24); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__25); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__26 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__26(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__26); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__27); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__28); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__29); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__30); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__31); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__32); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613____closed__33); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1613_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1613_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1613_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1673_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1673_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1673_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1751_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1751_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1751_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1821_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1821_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1821_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1891_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1891_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1891_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__2); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__5 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__5(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__5); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__8 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__8(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__8); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__9); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__10); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__11); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__12 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__12(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__12); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__14); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__17); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__19); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__20); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__21); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__23 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__23(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__23); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__26 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__26(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__26); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__27 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__27(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__27); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__29); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__30 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__30(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__30); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33 = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Core.c b/stage0/stdlib/Init/Core.c index 85ad0d0780af..8e5a5948a6ff 100644 --- a/stage0/stdlib/Init/Core.c +++ b/stage0/stdlib/Init/Core.c @@ -376,7 +376,6 @@ static lean_object* l_term___u2248_____closed__2; LEAN_EXPORT lean_object* l_instInhabitedMProd(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_PEmpty_elim___rarg(uint8_t); -LEAN_EXPORT lean_object* l_Subtype_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Core______macroRules__term___x5c____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Iff_elim(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Decidable_byCases(lean_object*, lean_object*); @@ -446,7 +445,6 @@ LEAN_EXPORT lean_object* l_Quot_hrecOn___rarg(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l___aux__Init__Core______macroRules__term___u2260____1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Core______macroRules__term_x7b_x7d__1___closed__3; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Subtype_instInhabited(lean_object*, lean_object*); static lean_object* l_term___u2283_____closed__6; static lean_object* l___aux__Init__Core______macroRules__term___x3c_x2d_x3e____1___closed__5; LEAN_EXPORT lean_object* l_Quotient_recOnSubsingleton_u2082(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -512,7 +510,6 @@ static lean_object* l___aux__Init__Core______macroRules__Lean__Parser__Tactic__t LEAN_EXPORT lean_object* l___aux__Init__Core______unexpand__HasEquiv__Equiv__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_instDecidableIte___rarg(uint8_t, uint8_t, uint8_t); static lean_object* l_term___u2283_____closed__4; -LEAN_EXPORT lean_object* l_Subtype_instInhabited___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_term_u2205; static lean_object* l___aux__Init__Core______macroRules__term___x3c_x2d_x3e____1___closed__1; static lean_object* l_term___u2260_____closed__2; @@ -6165,30 +6162,6 @@ x_2 = lean_alloc_closure((void*)(l_instInhabitedForInStep__1___rarg), 1, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Subtype_instInhabited___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_inc(x_1); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Subtype_instInhabited(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Subtype_instInhabited___rarg___boxed), 2, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Subtype_instInhabited___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Subtype_instInhabited___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} LEAN_EXPORT lean_object* l_Subtype_instDecidableEq___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { diff --git a/stage0/stdlib/Init/Data/Array/Basic.c b/stage0/stdlib/Init/Data/Array/Basic.c index c3c637cac635..a5f89192ba11 100644 --- a/stage0/stdlib/Init/Data/Array/Basic.c +++ b/stage0/stdlib/Init/Data/Array/Basic.c @@ -1261,7 +1261,7 @@ x_9 = l_Array_swapAt_x21___rarg___closed__2; x_10 = lean_string_append(x_8, x_9); x_11 = l_Array_swapAt_x21___rarg___closed__3; x_12 = l_Array_swapAt_x21___rarg___closed__4; -x_13 = lean_unsigned_to_nat(128u); +x_13 = lean_unsigned_to_nat(129u); x_14 = lean_unsigned_to_nat(4u); x_15 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_11, x_12, x_13, x_14, x_10); lean_dec(x_10); @@ -5554,7 +5554,7 @@ static lean_object* _init_l_Array_findSome_x21___rarg___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_swapAt_x21___rarg___closed__3; x_2 = l_Array_findSome_x21___rarg___closed__1; -x_3 = lean_unsigned_to_nat(449u); +x_3 = lean_unsigned_to_nat(452u); x_4 = lean_unsigned_to_nat(14u); x_5 = l_Array_findSome_x21___rarg___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -9684,7 +9684,7 @@ static lean_object* _init_l_Array_insertAt_x21___rarg___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_swapAt_x21___rarg___closed__3; x_2 = l_Array_insertAt_x21___rarg___closed__1; -x_3 = lean_unsigned_to_nat(774u); +x_3 = lean_unsigned_to_nat(782u); x_4 = lean_unsigned_to_nat(7u); x_5 = l_Array_insertAt_x21___rarg___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Init/Data/Array/Lemmas.c b/stage0/stdlib/Init/Data/Array/Lemmas.c index 525a7f799cab..a1184275b68f 100644 --- a/stage0/stdlib/Init/Data/Array/Lemmas.c +++ b/stage0/stdlib/Init/Data/Array/Lemmas.c @@ -13,31 +13,56 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_toListRev___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_foldlM_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_toListRev___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2(lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_drop_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_forIn_loop_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_forIn_loop_match__1_splitter(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_foldlM_loop_match__1_splitter(lean_object*); +LEAN_EXPORT lean_object* l_Array_instDecidableMemOfDecidableEq(lean_object*); LEAN_EXPORT lean_object* l_Array_toListRev___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter(lean_object*); +LEAN_EXPORT lean_object* l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_toListRev___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_forIn_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_toListRev___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_foldlM_loop_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_forIn_loop_match__1_splitter___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_drop_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_instDecidableMemOfDecidableEq___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_instDecidableMemOfDecidableEq___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_toListRev(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_drop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_foldlM_loop_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -283,6 +308,346 @@ lean_dec(x_1); return x_2; } } +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___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_eq(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_1, x_7); +x_9 = lean_apply_2(x_4, x_8, x_2); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_apply_1(x_3, x_2); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___rarg___boxed), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Init_Data_Array_Lemmas_0__Nat_fold_match__1_splitter___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___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_eq(x_1, x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_7 = lean_unsigned_to_nat(1u); +x_8 = lean_nat_sub(x_1, x_7); +x_9 = lean_apply_2(x_4, x_8, lean_box(0)); +return x_9; +} +else +{ +lean_object* x_10; +lean_dec(x_4); +x_10 = lean_apply_1(x_3, lean_box(0)); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___rarg___boxed), 4, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter___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_Init_Data_Array_Lemmas_0__Array_mapIdxM_map_match__1_splitter(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_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_dec(x_2); +lean_inc(x_3); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_2, x_4); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___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_Array_Lemmas_0__Array_findSomeRevM_x3f_find_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_dec(x_3); +lean_inc(x_2); +return x_2; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_3, x_4); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___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_Array_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___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___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___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_Array_Lemmas_0__Array_appendCore_loop_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___rarg(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); +lean_inc(x_1); +lean_inc(x_2); +x_8 = lean_apply_2(x_1, x_2, x_7); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +size_t x_10; size_t x_11; +x_10 = 1; +x_11 = lean_usize_add(x_4, x_10); +x_4 = x_11; +goto _start; +} +else +{ +uint8_t x_13; +lean_dec(x_2); +lean_dec(x_1); +x_13 = 1; +return x_13; +} +} +else +{ +uint8_t x_14; +lean_dec(x_2); +lean_dec(x_1); +x_14 = 0; +return x_14; +} +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_array_get_size(x_2); +x_5 = lean_unsigned_to_nat(0u); +x_6 = lean_nat_dec_lt(x_5, x_4); +if (x_6 == 0) +{ +uint8_t x_7; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_7 = 0; +return x_7; +} +else +{ +size_t x_8; size_t x_9; uint8_t x_10; +x_8 = 0; +x_9 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_10 = l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___rarg(x_1, x_3, x_2, x_8, x_9); +lean_dec(x_2); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT uint8_t l_Array_instDecidableMemOfDecidableEq___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg(x_1, x_3, x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_instDecidableMemOfDecidableEq(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_instDecidableMemOfDecidableEq___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Array_instDecidableMemOfDecidableEq___spec__2___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; 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_Array_instDecidableMemOfDecidableEq___spec__2___rarg(x_1, x_2, x_3, x_6, x_7); +lean_dec(x_3); +x_9 = lean_box(x_8); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Array_contains___at_Array_instDecidableMemOfDecidableEq___spec__1___rarg(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_instDecidableMemOfDecidableEq___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = l_Array_instDecidableMemOfDecidableEq___rarg(x_1, x_2, x_3); +x_5 = lean_box(x_4); +return x_5; +} +} lean_object* initialize_Init_Data_Nat_MinMax(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_Nat_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_Lemmas(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Init/Data/Array/Mem.c b/stage0/stdlib/Init/Data/Array/Mem.c index 149de5069784..6352820e7ef1 100644 --- a/stage0/stdlib/Init/Data/Array/Mem.c +++ b/stage0/stdlib/Init/Data/Array/Mem.c @@ -22,7 +22,9 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__25; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__2; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__24; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__38; @@ -43,6 +45,7 @@ static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__27; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__1; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__10; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); @@ -59,6 +62,7 @@ static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_tacticArray__get__dec___closed__5; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__20; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47; static lean_object* l_Array_tacticArray__get__dec___closed__3; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__2; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__14; @@ -74,13 +78,16 @@ static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__15; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__8; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__7; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__30; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__1; static lean_object* l_Array_tacticArray__get__dec___closed__6; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__12; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__21; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__8; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__41; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__4; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__36; @@ -91,6 +98,7 @@ static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__18; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__9; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__6; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__20; @@ -98,18 +106,22 @@ static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__37; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__4; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__45; static lean_object* l_Array_tacticArray__mem__dec___closed__2; LEAN_EXPORT lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_tacticArray__get__dec___closed__2; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__16; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__14; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__23; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__13; static lean_object* l_Array_tacticArray__mem__dec___closed__3; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39; static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__10; +static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__53; LEAN_EXPORT lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1(lean_object*, lean_object*, lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); static lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; @@ -322,7 +334,7 @@ static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("apply", 5); +x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } @@ -342,61 +354,56 @@ static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("sizeOf_get", 10); +x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__18() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; -x_2 = l_String_toSubstring_x27(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("withReducible", 13); +return x_1; } } static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__19() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; +x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__20() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_tacticArray__get__dec___closed__1; -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("with_reducible", 14); +return x_1; } } static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__21() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("apply", 5); +return x_1; } } static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__22() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; +x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__23() { @@ -497,52 +504,97 @@ return x_3; static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__33() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("paren", 5); -return x_1; -} -} -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__23; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__33; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(", 1); +x_1 = lean_mk_string_from_bytes("sizeOf_get", 10); return x_1; } } +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__36() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__37() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_tacticArray__get__dec___closed__1; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__38() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__38; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__40() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("ellipsis", 8); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__37() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__23; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__36; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__38() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__42() { _start: { lean_object* x_1; @@ -550,7 +602,7 @@ x_1 = lean_mk_string_from_bytes("..", 2); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43() { _start: { lean_object* x_1; @@ -558,7 +610,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__40() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44() { _start: { lean_object* x_1; @@ -566,7 +618,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__41() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__45() { _start: { lean_object* x_1; @@ -574,19 +626,19 @@ x_1 = lean_mk_string_from_bytes("simpArith", 9); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__42() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__41; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47() { _start: { lean_object* x_1; @@ -594,7 +646,7 @@ x_1 = lean_mk_string_from_bytes("simp_arith", 10); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; @@ -603,6 +655,67 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sizeOf_getElem", 14); +return x_1; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_tacticArray__get__dec___closed__1; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__53() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52; +x_3 = lean_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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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; +} +} LEAN_EXPORT lean_object* l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -621,7 +734,7 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; x_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); x_9 = 0; @@ -641,119 +754,188 @@ lean_inc(x_10); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_10); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__15; +x_17 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; lean_inc(x_10); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_10); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__19; -lean_inc(x_11); -lean_inc(x_12); -x_20 = l_Lean_addMacroScope(x_12, x_19, x_11); -x_21 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__18; -x_22 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__22; -lean_inc(x_10); -x_23 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_23, 0, x_10); -lean_ctor_set(x_23, 1, x_21); -lean_ctor_set(x_23, 2, x_20); -lean_ctor_set(x_23, 3, x_22); -x_24 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__16; -lean_inc(x_23); -lean_inc(x_18); +x_19 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__20; lean_inc(x_10); -x_25 = l_Lean_Syntax_node2(x_10, x_24, x_18, x_23); -x_26 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__7; +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__21; lean_inc(x_10); -x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); -x_28 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__14; +x_22 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_22, 0, x_10); +lean_ctor_set(x_22, 1, x_21); +x_23 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__30; +lean_inc(x_11); +lean_inc(x_12); +x_24 = l_Lean_addMacroScope(x_12, x_23, x_11); +x_25 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__27; +x_26 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__32; +lean_inc(x_10); +x_27 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_25); +lean_ctor_set(x_27, 2, x_24); +lean_ctor_set(x_27, 3, x_26); +x_28 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__36; +lean_inc(x_11); +lean_inc(x_12); +x_29 = l_Lean_addMacroScope(x_12, x_28, x_11); +x_30 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35; +x_31 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39; lean_inc(x_10); -x_29 = l_Lean_Syntax_node1(x_10, x_28, x_27); -x_30 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__12; +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_10); +lean_ctor_set(x_32, 1, x_30); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_31); +x_33 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__42; lean_inc(x_10); -x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); -x_32 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__9; -lean_inc(x_16); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_10); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__41; lean_inc(x_10); -x_33 = l_Lean_Syntax_node2(x_10, x_32, x_16, x_31); -x_34 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__30; -x_35 = l_Lean_addMacroScope(x_12, x_34, x_11); -x_36 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__27; -x_37 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__32; +x_36 = l_Lean_Syntax_node1(x_10, x_35, x_34); +x_37 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__7; lean_inc(x_10); -x_38 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_38, 0, x_10); -lean_ctor_set(x_38, 1, x_36); -lean_ctor_set(x_38, 2, x_35); -lean_ctor_set(x_38, 3, x_37); -x_39 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35; +x_38 = l_Lean_Syntax_node1(x_10, x_37, x_36); +x_39 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__25; +lean_inc(x_38); lean_inc(x_10); -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_10); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__38; +x_40 = l_Lean_Syntax_node2(x_10, x_39, x_32, x_38); +x_41 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43; lean_inc(x_10); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_10); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__37; +x_43 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__33; +lean_inc(x_42); +lean_inc(x_18); lean_inc(x_10); -x_44 = l_Lean_Syntax_node1(x_10, x_43, x_42); +x_44 = l_Lean_Syntax_node3(x_10, x_43, x_18, x_40, x_42); lean_inc(x_10); -x_45 = l_Lean_Syntax_node1(x_10, x_26, x_44); -x_46 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__25; +x_45 = l_Lean_Syntax_node1(x_10, x_37, x_44); +lean_inc(x_27); lean_inc(x_10); -x_47 = l_Lean_Syntax_node2(x_10, x_46, x_23, x_45); -x_48 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39; +x_46 = l_Lean_Syntax_node2(x_10, x_39, x_27, x_45); +x_47 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__22; +lean_inc(x_22); lean_inc(x_10); -x_49 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_49, 0, x_10); -lean_ctor_set(x_49, 1, x_48); -x_50 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34; +x_48 = l_Lean_Syntax_node2(x_10, x_47, x_22, x_46); lean_inc(x_10); -x_51 = l_Lean_Syntax_node3(x_10, x_50, x_40, x_47, x_49); +x_49 = l_Lean_Syntax_node1(x_10, x_37, x_48); +x_50 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__14; lean_inc(x_10); -x_52 = l_Lean_Syntax_node1(x_10, x_26, x_51); +x_51 = l_Lean_Syntax_node1(x_10, x_50, x_49); +x_52 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__12; lean_inc(x_10); -x_53 = l_Lean_Syntax_node2(x_10, x_46, x_38, x_52); +x_53 = l_Lean_Syntax_node1(x_10, x_52, x_51); +x_54 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__19; +lean_inc(x_20); lean_inc(x_10); -x_54 = l_Lean_Syntax_node2(x_10, x_24, x_18, x_53); -x_55 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__40; +x_55 = l_Lean_Syntax_node2(x_10, x_54, x_20, x_53); lean_inc(x_10); -x_56 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_56, 0, x_10); -lean_ctor_set(x_56, 1, x_55); -x_57 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43; +x_56 = l_Lean_Syntax_node1(x_10, x_37, x_55); lean_inc(x_10); -x_58 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_58, 0, x_10); -lean_ctor_set(x_58, 1, x_57); -x_59 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44; +x_57 = l_Lean_Syntax_node1(x_10, x_50, x_56); lean_inc(x_10); -x_60 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_60, 0, x_10); -lean_ctor_set(x_60, 1, x_26); -lean_ctor_set(x_60, 2, x_59); -x_61 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__42; -lean_inc_n(x_60, 4); +x_58 = l_Lean_Syntax_node1(x_10, x_52, x_57); +x_59 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__16; +lean_inc(x_42); +lean_inc(x_18); +lean_inc(x_10); +x_60 = l_Lean_Syntax_node3(x_10, x_59, x_18, x_58, x_42); +x_61 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44; +lean_inc(x_10); +x_62 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_62, 0, x_10); +lean_ctor_set(x_62, 1, x_61); +x_63 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47; +lean_inc(x_10); +x_64 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_64, 0, x_10); +lean_ctor_set(x_64, 1, x_63); +x_65 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48; +lean_inc(x_10); +x_66 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_66, 0, x_10); +lean_ctor_set(x_66, 1, x_37); +lean_ctor_set(x_66, 2, x_65); +x_67 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46; +lean_inc_n(x_66, 4); +lean_inc(x_10); +x_68 = l_Lean_Syntax_node6(x_10, x_67, x_64, x_66, x_66, x_66, x_66, x_66); +lean_inc(x_68); +lean_inc(x_62); +lean_inc(x_10); +x_69 = l_Lean_Syntax_node3(x_10, x_37, x_60, x_62, x_68); +lean_inc(x_10); +x_70 = l_Lean_Syntax_node1(x_10, x_50, x_69); +lean_inc(x_10); +x_71 = l_Lean_Syntax_node1(x_10, x_52, x_70); +x_72 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__9; +lean_inc(x_16); +lean_inc(x_10); +x_73 = l_Lean_Syntax_node2(x_10, x_72, x_16, x_71); +x_74 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51; +x_75 = l_Lean_addMacroScope(x_12, x_74, x_11); +x_76 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50; +x_77 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54; +lean_inc(x_10); +x_78 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_78, 0, x_10); +lean_ctor_set(x_78, 1, x_76); +lean_ctor_set(x_78, 2, x_75); +lean_ctor_set(x_78, 3, x_77); +lean_inc(x_10); +x_79 = l_Lean_Syntax_node2(x_10, x_39, x_78, x_38); +lean_inc(x_42); +lean_inc(x_18); +lean_inc(x_10); +x_80 = l_Lean_Syntax_node3(x_10, x_43, x_18, x_79, x_42); +lean_inc(x_10); +x_81 = l_Lean_Syntax_node1(x_10, x_37, x_80); lean_inc(x_10); -x_62 = l_Lean_Syntax_node6(x_10, x_61, x_58, x_60, x_60, x_60, x_60, x_60); +x_82 = l_Lean_Syntax_node2(x_10, x_39, x_27, x_81); lean_inc(x_10); -x_63 = l_Lean_Syntax_node3(x_10, x_26, x_54, x_56, x_62); +x_83 = l_Lean_Syntax_node2(x_10, x_47, x_22, x_82); lean_inc(x_10); -x_64 = l_Lean_Syntax_node1(x_10, x_28, x_63); +x_84 = l_Lean_Syntax_node1(x_10, x_37, x_83); lean_inc(x_10); -x_65 = l_Lean_Syntax_node1(x_10, x_30, x_64); +x_85 = l_Lean_Syntax_node1(x_10, x_50, x_84); lean_inc(x_10); -x_66 = l_Lean_Syntax_node2(x_10, x_32, x_16, x_65); +x_86 = l_Lean_Syntax_node1(x_10, x_52, x_85); lean_inc(x_10); -x_67 = l_Lean_Syntax_node2(x_10, x_26, x_33, x_66); -x_68 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__5; -x_69 = l_Lean_Syntax_node2(x_10, x_68, x_14, x_67); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_3); -return x_70; +x_87 = l_Lean_Syntax_node2(x_10, x_54, x_20, x_86); +lean_inc(x_10); +x_88 = l_Lean_Syntax_node1(x_10, x_37, x_87); +lean_inc(x_10); +x_89 = l_Lean_Syntax_node1(x_10, x_50, x_88); +lean_inc(x_10); +x_90 = l_Lean_Syntax_node1(x_10, x_52, x_89); +lean_inc(x_10); +x_91 = l_Lean_Syntax_node3(x_10, x_59, x_18, x_90, x_42); +lean_inc(x_10); +x_92 = l_Lean_Syntax_node3(x_10, x_37, x_91, x_62, x_68); +lean_inc(x_10); +x_93 = l_Lean_Syntax_node1(x_10, x_50, x_92); +lean_inc(x_10); +x_94 = l_Lean_Syntax_node1(x_10, x_52, x_93); +lean_inc(x_10); +x_95 = l_Lean_Syntax_node2(x_10, x_72, x_16, x_94); +lean_inc(x_10); +x_96 = l_Lean_Syntax_node2(x_10, x_37, x_73, x_95); +x_97 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__5; +x_98 = l_Lean_Syntax_node2(x_10, x_97, x_14, x_96); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_3); +return x_99; } } } @@ -975,24 +1157,48 @@ return x_5; static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__11() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__11; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__13() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("syntheticHole", 13); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__12() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__23; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__11; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__13() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__15() { _start: { lean_object* x_1; @@ -1000,7 +1206,7 @@ x_1 = lean_mk_string_from_bytes("\?", 1); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__14() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__16() { _start: { lean_object* x_1; @@ -1008,26 +1214,26 @@ x_1 = lean_mk_string_from_bytes("h", 1); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__15() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__14; +x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__16; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__16() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__14; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__19() { _start: { lean_object* x_1; @@ -1035,19 +1241,19 @@ x_1 = lean_mk_string_from_bytes("case'", 5); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__18() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__19() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__21() { _start: { lean_object* x_1; @@ -1055,19 +1261,19 @@ x_1 = lean_mk_string_from_bytes("caseArg", 7); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__20() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__2; x_3 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__3; -x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__19; +x_4 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__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___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__21() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23() { _start: { lean_object* x_1; @@ -1075,17 +1281,17 @@ x_1 = lean_mk_string_from_bytes("binderIdent", 11); return x_1; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__22() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__1; -x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__21; +x_2 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23() { +static lean_object* _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25() { _start: { lean_object* x_1; @@ -1111,7 +1317,7 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); x_9 = 0; @@ -1131,179 +1337,202 @@ lean_inc(x_10); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_10); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__15; +x_17 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__20; lean_inc(x_10); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_10); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__4; +x_19 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__21; +lean_inc(x_10); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__4; lean_inc(x_11); lean_inc(x_12); -x_20 = l_Lean_addMacroScope(x_12, x_19, x_11); -x_21 = lean_box(0); -x_22 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__2; -x_23 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__6; -lean_inc(x_10); -x_24 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_24, 0, x_10); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_20); -lean_ctor_set(x_24, 3, x_23); -x_25 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__16; -lean_inc(x_24); -lean_inc(x_18); -lean_inc(x_10); -x_26 = l_Lean_Syntax_node2(x_10, x_25, x_18, x_24); -x_27 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__40; -lean_inc(x_10); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_10); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__7; +x_22 = l_Lean_addMacroScope(x_12, x_21, x_11); +x_23 = lean_box(0); +x_24 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__2; +x_25 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__6; +lean_inc(x_10); +x_26 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_26, 0, x_10); +lean_ctor_set(x_26, 1, x_24); +lean_ctor_set(x_26, 2, x_22); +lean_ctor_set(x_26, 3, x_25); +x_27 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__22; +lean_inc(x_26); +lean_inc(x_20); +lean_inc(x_10); +x_28 = l_Lean_Syntax_node2(x_10, x_27, x_20, x_26); +x_29 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44; lean_inc(x_10); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_10); lean_ctor_set(x_30, 1, x_29); -x_31 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__8; +x_31 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__7; lean_inc(x_10); -x_32 = l_Lean_Syntax_node1(x_10, x_31, x_30); -x_33 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__9; +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_10); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__8; lean_inc(x_10); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_10); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__10; +x_34 = l_Lean_Syntax_node1(x_10, x_33, x_32); +x_35 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__9; lean_inc(x_10); -x_36 = l_Lean_Syntax_node1(x_10, x_35, x_34); -x_37 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__7; -lean_inc(x_32); -lean_inc(x_28); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_10); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__10; lean_inc(x_10); -x_38 = l_Lean_Syntax_node5(x_10, x_37, x_26, x_28, x_32, x_28, x_36); -x_39 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__14; +x_38 = l_Lean_Syntax_node1(x_10, x_37, x_36); +x_39 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__7; +lean_inc(x_34); +lean_inc(x_30); lean_inc(x_10); -x_40 = l_Lean_Syntax_node1(x_10, x_39, x_38); -x_41 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__12; +x_40 = l_Lean_Syntax_node5(x_10, x_39, x_28, x_30, x_34, x_30, x_38); +x_41 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__14; lean_inc(x_10); x_42 = l_Lean_Syntax_node1(x_10, x_41, x_40); -x_43 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__9; +x_43 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__12; +lean_inc(x_10); +x_44 = l_Lean_Syntax_node1(x_10, x_43, x_42); +x_45 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__19; +lean_inc(x_18); +lean_inc(x_10); +x_46 = l_Lean_Syntax_node2(x_10, x_45, x_18, x_44); +lean_inc(x_10); +x_47 = l_Lean_Syntax_node1(x_10, x_39, x_46); +lean_inc(x_10); +x_48 = l_Lean_Syntax_node1(x_10, x_41, x_47); +lean_inc(x_10); +x_49 = l_Lean_Syntax_node1(x_10, x_43, x_48); +x_50 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__9; lean_inc(x_16); lean_inc(x_10); -x_44 = l_Lean_Syntax_node2(x_10, x_43, x_16, x_42); -x_45 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__30; +x_51 = l_Lean_Syntax_node2(x_10, x_50, x_16, x_49); +x_52 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__30; lean_inc(x_11); lean_inc(x_12); -x_46 = l_Lean_addMacroScope(x_12, x_45, x_11); -x_47 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__27; -x_48 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__32; -lean_inc(x_10); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_10); -lean_ctor_set(x_49, 1, x_47); -lean_ctor_set(x_49, 2, x_46); -lean_ctor_set(x_49, 3, x_48); -x_50 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__35; -lean_inc(x_10); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_10); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__13; -lean_inc(x_10); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_10); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__16; -x_55 = l_Lean_addMacroScope(x_12, x_54, x_11); -x_56 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__15; -lean_inc(x_10); -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_10); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -lean_ctor_set(x_57, 3, x_21); -x_58 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__12; -lean_inc(x_57); -lean_inc(x_10); -x_59 = l_Lean_Syntax_node2(x_10, x_58, x_53, x_57); -lean_inc(x_10); -x_60 = l_Lean_Syntax_node1(x_10, x_37, x_59); -x_61 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__25; -lean_inc(x_10); -x_62 = l_Lean_Syntax_node2(x_10, x_61, x_24, x_60); -x_63 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__39; +x_53 = l_Lean_addMacroScope(x_12, x_52, x_11); +x_54 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__27; +x_55 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__12; lean_inc(x_10); -x_64 = lean_alloc_ctor(2, 2, 0); +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_10); +lean_ctor_set(x_56, 1, x_54); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_55); +x_57 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__17; +lean_inc(x_10); +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_10); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__15; +lean_inc(x_10); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_10); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__18; +x_62 = l_Lean_addMacroScope(x_12, x_61, x_11); +x_63 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17; +lean_inc(x_10); +x_64 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_64, 0, x_10); lean_ctor_set(x_64, 1, x_63); -x_65 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__34; -lean_inc(x_10); -x_66 = l_Lean_Syntax_node3(x_10, x_65, x_51, x_62, x_64); +lean_ctor_set(x_64, 2, x_62); +lean_ctor_set(x_64, 3, x_23); +x_65 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__14; +lean_inc(x_64); lean_inc(x_10); -x_67 = l_Lean_Syntax_node1(x_10, x_37, x_66); +x_66 = l_Lean_Syntax_node2(x_10, x_65, x_60, x_64); lean_inc(x_10); -x_68 = l_Lean_Syntax_node2(x_10, x_61, x_49, x_67); +x_67 = l_Lean_Syntax_node1(x_10, x_39, x_66); +x_68 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__25; lean_inc(x_10); -x_69 = l_Lean_Syntax_node2(x_10, x_25, x_18, x_68); -x_70 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44; +x_69 = l_Lean_Syntax_node2(x_10, x_68, x_26, x_67); +x_70 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43; lean_inc(x_10); -x_71 = lean_alloc_ctor(1, 3, 0); +x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_10); -lean_ctor_set(x_71, 1, x_37); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__17; +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__33; lean_inc(x_10); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_10); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__22; +x_73 = l_Lean_Syntax_node3(x_10, x_72, x_58, x_69, x_71); lean_inc(x_10); -x_75 = l_Lean_Syntax_node1(x_10, x_74, x_57); -x_76 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__20; -lean_inc(x_71); +x_74 = l_Lean_Syntax_node1(x_10, x_39, x_73); lean_inc(x_10); -x_77 = l_Lean_Syntax_node2(x_10, x_76, x_75, x_71); +x_75 = l_Lean_Syntax_node2(x_10, x_68, x_56, x_74); lean_inc(x_10); -x_78 = l_Lean_Syntax_node1(x_10, x_37, x_77); -x_79 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23; +x_76 = l_Lean_Syntax_node2(x_10, x_27, x_20, x_75); +x_77 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48; +lean_inc(x_10); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_10); +lean_ctor_set(x_78, 1, x_39); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__19; lean_inc(x_10); x_80 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_80, 0, x_10); lean_ctor_set(x_80, 1, x_79); +x_81 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24; lean_inc(x_10); -x_81 = l_Lean_Syntax_node1(x_10, x_37, x_32); -lean_inc(x_10); -x_82 = l_Lean_Syntax_node1(x_10, x_39, x_81); +x_82 = l_Lean_Syntax_node1(x_10, x_81, x_64); +x_83 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__22; +lean_inc(x_78); lean_inc(x_10); -x_83 = l_Lean_Syntax_node1(x_10, x_41, x_82); -x_84 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__18; +x_84 = l_Lean_Syntax_node2(x_10, x_83, x_82, x_78); lean_inc(x_10); -x_85 = l_Lean_Syntax_node4(x_10, x_84, x_73, x_78, x_80, x_83); -x_86 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43; +x_85 = l_Lean_Syntax_node1(x_10, x_39, x_84); +x_86 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25; lean_inc(x_10); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_10); lean_ctor_set(x_87, 1, x_86); -x_88 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__42; -lean_inc_n(x_71, 5); lean_inc(x_10); -x_89 = l_Lean_Syntax_node6(x_10, x_88, x_87, x_71, x_71, x_71, x_71, x_71); -lean_inc(x_71); +x_88 = l_Lean_Syntax_node1(x_10, x_39, x_34); +lean_inc(x_10); +x_89 = l_Lean_Syntax_node1(x_10, x_41, x_88); +lean_inc(x_10); +x_90 = l_Lean_Syntax_node1(x_10, x_43, x_89); +x_91 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__20; +lean_inc(x_10); +x_92 = l_Lean_Syntax_node4(x_10, x_91, x_80, x_85, x_87, x_90); +lean_inc(x_78); +lean_inc(x_10); +x_93 = l_Lean_Syntax_node3(x_10, x_39, x_76, x_78, x_92); +lean_inc(x_10); +x_94 = l_Lean_Syntax_node1(x_10, x_41, x_93); +lean_inc(x_10); +x_95 = l_Lean_Syntax_node1(x_10, x_43, x_94); +lean_inc(x_10); +x_96 = l_Lean_Syntax_node2(x_10, x_45, x_18, x_95); +x_97 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47; +lean_inc(x_10); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_10); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46; +lean_inc_n(x_78, 5); +lean_inc(x_10); +x_100 = l_Lean_Syntax_node6(x_10, x_99, x_98, x_78, x_78, x_78, x_78, x_78); lean_inc(x_10); -x_90 = l_Lean_Syntax_node5(x_10, x_37, x_69, x_71, x_85, x_71, x_89); +x_101 = l_Lean_Syntax_node3(x_10, x_39, x_96, x_78, x_100); lean_inc(x_10); -x_91 = l_Lean_Syntax_node1(x_10, x_39, x_90); +x_102 = l_Lean_Syntax_node1(x_10, x_41, x_101); lean_inc(x_10); -x_92 = l_Lean_Syntax_node1(x_10, x_41, x_91); +x_103 = l_Lean_Syntax_node1(x_10, x_43, x_102); lean_inc(x_10); -x_93 = l_Lean_Syntax_node2(x_10, x_43, x_16, x_92); +x_104 = l_Lean_Syntax_node2(x_10, x_50, x_16, x_103); lean_inc(x_10); -x_94 = l_Lean_Syntax_node2(x_10, x_37, x_44, x_93); -x_95 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__5; -x_96 = l_Lean_Syntax_node2(x_10, x_95, x_14, x_94); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_3); -return x_97; +x_105 = l_Lean_Syntax_node2(x_10, x_39, x_51, x_104); +x_106 = l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__5; +x_107 = l_Lean_Syntax_node2(x_10, x_106, x_14, x_105); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_3); +return x_108; } } } @@ -1464,6 +1693,26 @@ l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__ lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__43); l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44(); lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__44); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__45 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__45(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__45); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__46); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__47); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__48); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__49); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__50); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__51); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__52); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__53 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__53(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__53); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__get__dec__1___closed__54); l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__1 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__1(); lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__1); l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__2 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__tacticDecreasing__trivial__1___closed__2(); @@ -1526,6 +1775,10 @@ l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__ lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__22); l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23(); lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__23); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__24); +l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25 = _init_l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25(); +lean_mark_persistent(l_Array___aux__Init__Data__Array__Mem______macroRules__Array__tacticArray__mem__dec__1___closed__25); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Data/List.c b/stage0/stdlib/Init/Data/List.c index 5a3e7e3066a8..9e5337499293 100644 --- a/stage0/stdlib/Init/Data/List.c +++ b/stage0/stdlib/Init/Data/List.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.List -// Imports: Init.Data.List.Basic Init.Data.List.BasicAux Init.Data.List.Control Init.Data.List.Lemmas Init.Data.List.Impl +// Imports: Init.Data.List.Basic Init.Data.List.BasicAux Init.Data.List.Control Init.Data.List.Lemmas Init.Data.List.Impl Init.Data.List.TakeDrop #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -18,6 +18,7 @@ lean_object* initialize_Init_Data_List_BasicAux(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_Control(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_Impl(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_List_TakeDrop(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Init_Data_List(uint8_t builtin, lean_object* w) { lean_object * res; @@ -38,6 +39,9 @@ lean_dec_ref(res); res = initialize_Init_Data_List_Impl(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Init_Data_List_TakeDrop(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/List/BasicAux.c b/stage0/stdlib/Init/Data/List/BasicAux.c index 1a98faa07f1e..3ec69f89a146 100644 --- a/stage0/stdlib/Init/Data/List/BasicAux.c +++ b/stage0/stdlib/Init/Data/List/BasicAux.c @@ -61,6 +61,7 @@ LEAN_EXPORT lean_object* l_List_partitionM_go___rarg(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; LEAN_EXPORT lean_object* l_List_head___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61; LEAN_EXPORT lean_object* l_List_getD(lean_object*); static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__18; @@ -105,6 +106,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__40; +static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60; static lean_object* l_List_get_x21___rarg___closed__1; LEAN_EXPORT lean_object* l_List_get_x21___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_BasicAux_0__List_mapMonoMImp___at_List_mapMono___spec__1___rarg(lean_object*, lean_object*); @@ -129,6 +131,7 @@ static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__ta static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; LEAN_EXPORT lean_object* l_List_get_x21(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__59; static lean_object* l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; LEAN_EXPORT lean_object* l_List_headD___rarg(lean_object*, lean_object*); lean_object* l_List_drop___rarg(lean_object*, lean_object*); @@ -1234,7 +1237,7 @@ static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRul _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("apply", 5); +x_1 = lean_mk_string_from_bytes("withReducible", 13); return x_1; } } @@ -1254,64 +1257,92 @@ static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRul _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("sizeOf_lt_of_mem", 16); +x_1 = lean_mk_string_from_bytes("with_reducible", 14); return x_1; } } static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__18() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("apply", 5); +return x_1; +} +} +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; +x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("sizeOf_lt_of_mem", 16); +return x_1; +} +} +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; +x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__19() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_List_tacticSizeOf__list__dec___closed__1; -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__20; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__22() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__23() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__26() { _start: { lean_object* x_1; @@ -1319,7 +1350,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__27() { _start: { lean_object* x_1; @@ -1327,19 +1358,19 @@ x_1 = lean_mk_string_from_bytes("assumption", 10); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__26() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29() { _start: { lean_object* x_1; @@ -1347,19 +1378,19 @@ x_1 = lean_mk_string_from_bytes("done", 4); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__27() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__26; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31() { _start: { lean_object* x_1; @@ -1367,7 +1398,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__32() { _start: { lean_object* x_1; @@ -1375,19 +1406,19 @@ x_1 = lean_mk_string_from_bytes("app", 3); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__30() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; -x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29; +x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34() { _start: { lean_object* x_1; @@ -1395,16 +1426,16 @@ x_1 = lean_mk_string_from_bytes("Nat.lt_trans", 12); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__32() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; +x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__33() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__36() { _start: { lean_object* x_1; @@ -1412,7 +1443,7 @@ x_1 = lean_mk_string_from_bytes("Nat", 3); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37() { _start: { lean_object* x_1; @@ -1420,41 +1451,41 @@ x_1 = lean_mk_string_from_bytes("lt_trans", 8); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__33; -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__34; +x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__36; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__36() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__40() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__36; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__41() { _start: { lean_object* x_1; @@ -1462,19 +1493,19 @@ x_1 = lean_mk_string_from_bytes("paren", 5); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__39() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; -x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38; +x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__40() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__43() { _start: { lean_object* x_1; @@ -1482,7 +1513,7 @@ x_1 = lean_mk_string_from_bytes("(", 1); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__41() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__44() { _start: { lean_object* x_1; @@ -1490,19 +1521,19 @@ x_1 = lean_mk_string_from_bytes("syntheticHole", 13); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__42() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; -x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__41; +x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__31; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__44; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__43() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__46() { _start: { lean_object* x_1; @@ -1510,7 +1541,7 @@ x_1 = lean_mk_string_from_bytes("\?", 1); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__44() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47() { _start: { lean_object* x_1; @@ -1518,26 +1549,26 @@ x_1 = lean_mk_string_from_bytes("h", 1); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__44; +x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__46() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__44; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__50() { _start: { lean_object* x_1; @@ -1545,7 +1576,7 @@ x_1 = lean_mk_string_from_bytes(")", 1); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51() { _start: { lean_object* x_1; lean_object* x_2; @@ -1554,7 +1585,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__52() { _start: { lean_object* x_1; @@ -1562,19 +1593,19 @@ x_1 = lean_mk_string_from_bytes("case'", 5); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__50() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54() { _start: { lean_object* x_1; @@ -1582,19 +1613,19 @@ x_1 = lean_mk_string_from_bytes("caseArg", 7); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__52() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__53() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56() { _start: { lean_object* x_1; @@ -1602,17 +1633,17 @@ x_1 = lean_mk_string_from_bytes("binderIdent", 11); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; -x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__53; +x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__55() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58() { _start: { lean_object* x_1; @@ -1620,7 +1651,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__59() { _start: { lean_object* x_1; @@ -1628,19 +1659,19 @@ x_1 = lean_mk_string_from_bytes("simpArith", 9); return x_1; } } -static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__1; x_2 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__2; x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__3; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__56; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__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_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58() { +static lean_object* _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61() { _start: { lean_object* x_1; @@ -1666,7 +1697,7 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; x_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); x_9 = 0; @@ -1686,179 +1717,202 @@ lean_inc(x_10); x_16 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_16, 0, x_10); lean_ctor_set(x_16, 1, x_15); -x_17 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__15; +x_17 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__17; lean_inc(x_10); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_10); lean_ctor_set(x_18, 1, x_17); -x_19 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__19; +x_19 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__18; +lean_inc(x_10); +x_20 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_20, 0, x_10); +lean_ctor_set(x_20, 1, x_19); +x_21 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__22; lean_inc(x_11); lean_inc(x_12); -x_20 = l_Lean_addMacroScope(x_12, x_19, x_11); -x_21 = lean_box(0); -x_22 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__18; -x_23 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__22; -lean_inc(x_10); -x_24 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_24, 0, x_10); -lean_ctor_set(x_24, 1, x_22); -lean_ctor_set(x_24, 2, x_20); -lean_ctor_set(x_24, 3, x_23); -x_25 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__16; -lean_inc(x_24); -lean_inc(x_18); +x_22 = l_Lean_addMacroScope(x_12, x_21, x_11); +x_23 = lean_box(0); +x_24 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__21; +x_25 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25; lean_inc(x_10); -x_26 = l_Lean_Syntax_node2(x_10, x_25, x_18, x_24); -x_27 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__23; +x_26 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_26, 0, x_10); +lean_ctor_set(x_26, 1, x_24); +lean_ctor_set(x_26, 2, x_22); +lean_ctor_set(x_26, 3, x_25); +x_27 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__19; +lean_inc(x_26); +lean_inc(x_20); lean_inc(x_10); -x_28 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_28, 0, x_10); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__24; +x_28 = l_Lean_Syntax_node2(x_10, x_27, x_20, x_26); +x_29 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__26; lean_inc(x_10); x_30 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_30, 0, x_10); lean_ctor_set(x_30, 1, x_29); -x_31 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__25; +x_31 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__27; lean_inc(x_10); -x_32 = l_Lean_Syntax_node1(x_10, x_31, x_30); -x_33 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__26; +x_32 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_32, 0, x_10); +lean_ctor_set(x_32, 1, x_31); +x_33 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__28; lean_inc(x_10); -x_34 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_34, 0, x_10); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__27; +x_34 = l_Lean_Syntax_node1(x_10, x_33, x_32); +x_35 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__29; lean_inc(x_10); -x_36 = l_Lean_Syntax_node1(x_10, x_35, x_34); -x_37 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__7; -lean_inc(x_32); -lean_inc(x_28); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_10); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__30; lean_inc(x_10); -x_38 = l_Lean_Syntax_node5(x_10, x_37, x_26, x_28, x_32, x_28, x_36); -x_39 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__14; +x_38 = l_Lean_Syntax_node1(x_10, x_37, x_36); +x_39 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__7; +lean_inc(x_34); +lean_inc(x_30); lean_inc(x_10); -x_40 = l_Lean_Syntax_node1(x_10, x_39, x_38); -x_41 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; +x_40 = l_Lean_Syntax_node5(x_10, x_39, x_28, x_30, x_34, x_30, x_38); +x_41 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__14; lean_inc(x_10); x_42 = l_Lean_Syntax_node1(x_10, x_41, x_40); -x_43 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__9; -lean_inc(x_16); +x_43 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__12; lean_inc(x_10); -x_44 = l_Lean_Syntax_node2(x_10, x_43, x_16, x_42); -x_45 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; -lean_inc(x_11); -lean_inc(x_12); -x_46 = l_Lean_addMacroScope(x_12, x_45, x_11); -x_47 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__32; -x_48 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__37; +x_44 = l_Lean_Syntax_node1(x_10, x_43, x_42); +x_45 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__16; +lean_inc(x_18); lean_inc(x_10); -x_49 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_49, 0, x_10); -lean_ctor_set(x_49, 1, x_47); -lean_ctor_set(x_49, 2, x_46); -lean_ctor_set(x_49, 3, x_48); -x_50 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__40; +x_46 = l_Lean_Syntax_node2(x_10, x_45, x_18, x_44); lean_inc(x_10); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_10); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__43; +x_47 = l_Lean_Syntax_node1(x_10, x_39, x_46); lean_inc(x_10); -x_53 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_53, 0, x_10); -lean_ctor_set(x_53, 1, x_52); -x_54 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__46; -x_55 = l_Lean_addMacroScope(x_12, x_54, x_11); -x_56 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45; +x_48 = l_Lean_Syntax_node1(x_10, x_41, x_47); lean_inc(x_10); -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_10); -lean_ctor_set(x_57, 1, x_56); -lean_ctor_set(x_57, 2, x_55); -lean_ctor_set(x_57, 3, x_21); -x_58 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__42; -lean_inc(x_57); +x_49 = l_Lean_Syntax_node1(x_10, x_43, x_48); +x_50 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__9; +lean_inc(x_16); lean_inc(x_10); -x_59 = l_Lean_Syntax_node2(x_10, x_58, x_53, x_57); +x_51 = l_Lean_Syntax_node2(x_10, x_50, x_16, x_49); +x_52 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__38; +lean_inc(x_11); +lean_inc(x_12); +x_53 = l_Lean_addMacroScope(x_12, x_52, x_11); +x_54 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__35; +x_55 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__40; lean_inc(x_10); -x_60 = l_Lean_Syntax_node1(x_10, x_37, x_59); -x_61 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__30; +x_56 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_56, 0, x_10); +lean_ctor_set(x_56, 1, x_54); +lean_ctor_set(x_56, 2, x_53); +lean_ctor_set(x_56, 3, x_55); +x_57 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__43; lean_inc(x_10); -x_62 = l_Lean_Syntax_node2(x_10, x_61, x_24, x_60); -x_63 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__47; +x_58 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_58, 0, x_10); +lean_ctor_set(x_58, 1, x_57); +x_59 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__46; lean_inc(x_10); -x_64 = lean_alloc_ctor(2, 2, 0); +x_60 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_60, 0, x_10); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49; +x_62 = l_Lean_addMacroScope(x_12, x_61, x_11); +x_63 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48; +lean_inc(x_10); +x_64 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_64, 0, x_10); lean_ctor_set(x_64, 1, x_63); -x_65 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__39; -lean_inc(x_10); -x_66 = l_Lean_Syntax_node3(x_10, x_65, x_51, x_62, x_64); +lean_ctor_set(x_64, 2, x_62); +lean_ctor_set(x_64, 3, x_23); +x_65 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__45; +lean_inc(x_64); lean_inc(x_10); -x_67 = l_Lean_Syntax_node1(x_10, x_37, x_66); +x_66 = l_Lean_Syntax_node2(x_10, x_65, x_60, x_64); lean_inc(x_10); -x_68 = l_Lean_Syntax_node2(x_10, x_61, x_49, x_67); +x_67 = l_Lean_Syntax_node1(x_10, x_39, x_66); +x_68 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__33; lean_inc(x_10); -x_69 = l_Lean_Syntax_node2(x_10, x_25, x_18, x_68); -x_70 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48; +x_69 = l_Lean_Syntax_node2(x_10, x_68, x_26, x_67); +x_70 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__50; lean_inc(x_10); -x_71 = lean_alloc_ctor(1, 3, 0); +x_71 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_71, 0, x_10); -lean_ctor_set(x_71, 1, x_37); -lean_ctor_set(x_71, 2, x_70); -x_72 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__49; +lean_ctor_set(x_71, 1, x_70); +x_72 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__42; lean_inc(x_10); -x_73 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_73, 0, x_10); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__54; +x_73 = l_Lean_Syntax_node3(x_10, x_72, x_58, x_69, x_71); lean_inc(x_10); -x_75 = l_Lean_Syntax_node1(x_10, x_74, x_57); -x_76 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__52; -lean_inc(x_71); +x_74 = l_Lean_Syntax_node1(x_10, x_39, x_73); lean_inc(x_10); -x_77 = l_Lean_Syntax_node2(x_10, x_76, x_75, x_71); +x_75 = l_Lean_Syntax_node2(x_10, x_68, x_56, x_74); lean_inc(x_10); -x_78 = l_Lean_Syntax_node1(x_10, x_37, x_77); -x_79 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__55; +x_76 = l_Lean_Syntax_node2(x_10, x_27, x_20, x_75); +x_77 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; +lean_inc(x_10); +x_78 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_78, 0, x_10); +lean_ctor_set(x_78, 1, x_39); +lean_ctor_set(x_78, 2, x_77); +x_79 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__52; lean_inc(x_10); x_80 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_80, 0, x_10); lean_ctor_set(x_80, 1, x_79); +x_81 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57; lean_inc(x_10); -x_81 = l_Lean_Syntax_node1(x_10, x_37, x_32); -lean_inc(x_10); -x_82 = l_Lean_Syntax_node1(x_10, x_39, x_81); +x_82 = l_Lean_Syntax_node1(x_10, x_81, x_64); +x_83 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__55; +lean_inc(x_78); lean_inc(x_10); -x_83 = l_Lean_Syntax_node1(x_10, x_41, x_82); -x_84 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__50; +x_84 = l_Lean_Syntax_node2(x_10, x_83, x_82, x_78); lean_inc(x_10); -x_85 = l_Lean_Syntax_node4(x_10, x_84, x_73, x_78, x_80, x_83); +x_85 = l_Lean_Syntax_node1(x_10, x_39, x_84); x_86 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58; lean_inc(x_10); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_10); lean_ctor_set(x_87, 1, x_86); -x_88 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57; -lean_inc_n(x_71, 5); lean_inc(x_10); -x_89 = l_Lean_Syntax_node6(x_10, x_88, x_87, x_71, x_71, x_71, x_71, x_71); -lean_inc(x_71); +x_88 = l_Lean_Syntax_node1(x_10, x_39, x_34); +lean_inc(x_10); +x_89 = l_Lean_Syntax_node1(x_10, x_41, x_88); +lean_inc(x_10); +x_90 = l_Lean_Syntax_node1(x_10, x_43, x_89); +x_91 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__53; lean_inc(x_10); -x_90 = l_Lean_Syntax_node5(x_10, x_37, x_69, x_71, x_85, x_71, x_89); +x_92 = l_Lean_Syntax_node4(x_10, x_91, x_80, x_85, x_87, x_90); +lean_inc(x_78); lean_inc(x_10); -x_91 = l_Lean_Syntax_node1(x_10, x_39, x_90); +x_93 = l_Lean_Syntax_node3(x_10, x_39, x_76, x_78, x_92); lean_inc(x_10); -x_92 = l_Lean_Syntax_node1(x_10, x_41, x_91); +x_94 = l_Lean_Syntax_node1(x_10, x_41, x_93); lean_inc(x_10); -x_93 = l_Lean_Syntax_node2(x_10, x_43, x_16, x_92); +x_95 = l_Lean_Syntax_node1(x_10, x_43, x_94); lean_inc(x_10); -x_94 = l_Lean_Syntax_node2(x_10, x_37, x_44, x_93); -x_95 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__5; -x_96 = l_Lean_Syntax_node2(x_10, x_95, x_14, x_94); -x_97 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_3); -return x_97; +x_96 = l_Lean_Syntax_node2(x_10, x_45, x_18, x_95); +x_97 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61; +lean_inc(x_10); +x_98 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_98, 0, x_10); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60; +lean_inc_n(x_78, 5); +lean_inc(x_10); +x_100 = l_Lean_Syntax_node6(x_10, x_99, x_98, x_78, x_78, x_78, x_78, x_78); +lean_inc(x_10); +x_101 = l_Lean_Syntax_node3(x_10, x_39, x_96, x_78, x_100); +lean_inc(x_10); +x_102 = l_Lean_Syntax_node1(x_10, x_41, x_101); +lean_inc(x_10); +x_103 = l_Lean_Syntax_node1(x_10, x_43, x_102); +lean_inc(x_10); +x_104 = l_Lean_Syntax_node2(x_10, x_50, x_16, x_103); +lean_inc(x_10); +x_105 = l_Lean_Syntax_node2(x_10, x_39, x_51, x_104); +x_106 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__5; +x_107 = l_Lean_Syntax_node2(x_10, x_106, x_14, x_105); +x_108 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_3); +return x_108; } } } @@ -2277,7 +2331,7 @@ LEAN_EXPORT lean_object* l_List_partitionM___rarg(lean_object* x_1, lean_object* _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48; +x_4 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; x_5 = l_List_partitionM_go___rarg(x_1, x_2, x_3, x_4, x_4); return x_5; } @@ -2351,7 +2405,7 @@ LEAN_EXPORT lean_object* l_List_partitionMap___rarg(lean_object* x_1, lean_objec _start: { lean_object* x_3; lean_object* x_4; -x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__48; +x_3 = l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__51; x_4 = l_List_partitionMap_go___rarg(x_1, x_2, x_3, x_3); return x_4; } @@ -2529,6 +2583,12 @@ l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__li lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__57); l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58(); lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__58); +l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__59 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__59(); +lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__59); +l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60(); +lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__60); +l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61(); +lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__List__tacticSizeOf__list__dec__1___closed__61); l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1(); lean_mark_persistent(l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__1); l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__2 = _init_l_List___aux__Init__Data__List__BasicAux______macroRules__tacticDecreasing__trivial__1___closed__2(); diff --git a/stage0/stdlib/Init/Data/List/Impl.c b/stage0/stdlib/Init/Data/List/Impl.c index f22de306e143..a4de1aed2faf 100644 --- a/stage0/stdlib/Init/Data/List/Impl.c +++ b/stage0/stdlib/Init/Data/List/Impl.c @@ -108,14 +108,12 @@ LEAN_EXPORT lean_object* l_List_takeWhileTR(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_intersperseTR___spec__3___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_setTR_go_match__1_splitter(lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_erase_match__2_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_enumFromTR___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_replace_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_intercalateTR_go_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_takeTR_go(lean_object*); LEAN_EXPORT lean_object* l_List_foldrTR___at_List_unzipTR___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldrTR___at_List_unzipTR___spec__1___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_eraseTR_go___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_enumFromTR___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_intersperseTR_match__1_splitter(lean_object*, lean_object*); @@ -125,7 +123,6 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_takeTR_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_enumFrom_match__1_splitter(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_erase_match__2_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_zipWithTR_go(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterMapTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_replace_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -145,10 +142,8 @@ LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_filterMapTR_go_ LEAN_EXPORT lean_object* l_List_replaceTR_go(lean_object*); LEAN_EXPORT lean_object* l_List_setTR___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_setTR_go___spec__1(lean_object*); LEAN_EXPORT lean_object* l_List_foldrTR___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter(lean_object*, lean_object*); lean_object* l_Array_foldrMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_List_intercalateTR_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -728,38 +723,6 @@ x_2 = lean_alloc_closure((void*)(l_List_eraseTR___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_erase_match__2_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_5; -lean_dec(x_4); -x_5 = lean_apply_1(x_3, x_2); -return x_5; -} -else -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_1, 1); -lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_apply_3(x_4, x_6, x_7, x_2); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_erase_match__2_splitter(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Impl_0__List_erase_match__2_splitter___rarg), 4, 0); -return x_3; -} -} LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_eraseTR_go_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -1276,45 +1239,6 @@ x_3 = lean_alloc_closure((void*)(l_List_filterMapTR___rarg), 2, 0); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_dec(x_3); -lean_inc(x_2); -return x_2; -} -else -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_1, 1); -lean_inc(x_5); -lean_dec(x_1); -x_6 = lean_apply_2(x_3, x_4, x_5); -return x_6; -} -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter___rarg___boxed), 3, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_isEmpty_match__1_splitter___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_List_Impl_0__List_isEmpty_match__1_splitter___rarg(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} -} LEAN_EXPORT lean_object* l___private_Init_Data_List_Impl_0__List_filterMapTR_go_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { diff --git a/stage0/stdlib/Init/Data/List/Lemmas.c b/stage0/stdlib/Init/Data/List/Lemmas.c index ec67b0c43167..8c03e3d98c5e 100644 --- a/stage0/stdlib/Init/Data/List/Lemmas.c +++ b/stage0/stdlib/Init/Data/List/Lemmas.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Init.Data.List.Lemmas -// Imports: Init.Data.List.BasicAux Init.Data.List.Control Init.PropLemmas Init.Control.Lawful.Basic Init.Hints +// Imports: Init.Data.Bool Init.Data.Option.Lemmas Init.Data.List.BasicAux Init.Data.List.Control Init.PropLemmas Init.Control.Lawful.Basic Init.Hints #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -20,17 +20,26 @@ LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_zipWith__get_ LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_isPrefixOf_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_join_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_isPrefixOf_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__2_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_join_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findM_x3f_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_getLast_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg(uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapM_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_x27___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__2_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldlM_match__1_splitter(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_x27(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapA_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_zipWithAll__get_x3f_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldlM_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapM_x27_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapM_x27_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_isEmpty_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterAuxM_match__1_splitter(lean_object*, lean_object*); @@ -38,19 +47,26 @@ LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__ LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_isEmpty_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findM_x3f_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapM_loop_match__1_splitter(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_zipWithAll__get_x3f_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_x27___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_isEmpty_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapA_match__1_splitter(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findM_x3f_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_join_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_get_x21_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_get_x21_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapA_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_foldl_match__1_splitter(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_zipWith_match__1_splitter(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg(uint8_t, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_mapM_x27_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_zipWith__get_x3f_match__1_splitter(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_x27___rarg(lean_object*, lean_object*, lean_object*); @@ -742,6 +758,226 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_inc(x_3); +return x_3; +} +else +{ +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Init_Data_List_Lemmas_0__List_erase_match__1_splitter___rarg(x_4, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__2_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; +lean_dec(x_4); +x_5 = lean_apply_1(x_3, x_2); +return x_5; +} +else +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_1, 1); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_apply_3(x_4, x_6, x_7, x_2); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_erase_match__2_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_erase_match__2_splitter___rarg), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_dec(x_4); +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_apply_2(x_3, x_5, x_6); +return x_7; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_3); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); +lean_dec(x_1); +x_10 = lean_ctor_get(x_2, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_2, 1); +lean_inc(x_11); +lean_dec(x_2); +x_12 = lean_apply_4(x_4, x_8, x_9, x_10, x_11); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_partition_loop_match__1_splitter___rarg), 4, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_dec(x_3); +lean_inc(x_2); +return x_2; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_3, x_4); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_filterMap__cons_match__1_splitter___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_List_Lemmas_0__List_filterMap__cons_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg(uint8_t x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (x_1 == 0) +{ +lean_inc(x_3); +return x_3; +} +else +{ +lean_inc(x_2); +return x_2; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = l___private_Init_Data_List_Lemmas_0__List_find_x3f__cons_match__1_splitter___rarg(x_4, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_dec(x_2); +lean_inc(x_3); +return x_3; +} +else +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_apply_1(x_2, x_4); +return x_5; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___rarg___boxed), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___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_List_Lemmas_0__List_findSome_x3f__cons_match__1_splitter___rarg(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +lean_object* initialize_Init_Data_Bool(uint8_t builtin, lean_object*); +lean_object* initialize_Init_Data_Option_Lemmas(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_BasicAux(uint8_t builtin, lean_object*); lean_object* initialize_Init_Data_List_Control(uint8_t builtin, lean_object*); lean_object* initialize_Init_PropLemmas(uint8_t builtin, lean_object*); @@ -752,6 +988,12 @@ LEAN_EXPORT lean_object* initialize_Init_Data_List_Lemmas(uint8_t builtin, lean_ lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _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_Option_Lemmas(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Init_Data_List_BasicAux(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/List/TakeDrop.c b/stage0/stdlib/Init/Data/List/TakeDrop.c new file mode 100644 index 000000000000..49eb4ea42ded --- /dev/null +++ b/stage0/stdlib/Init/Data/List/TakeDrop.c @@ -0,0 +1,97 @@ +// Lean compiler output +// Module: Init.Data.List.TakeDrop +// Imports: Init.Data.List.Lemmas 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_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___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; uint8_t x_7; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_eq(x_1, x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +lean_dec(x_3); +x_8 = lean_unsigned_to_nat(1u); +x_9 = lean_nat_sub(x_1, x_8); +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_10; +lean_dec(x_5); +x_10 = lean_apply_1(x_4, x_9); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_dec(x_4); +x_11 = lean_ctor_get(x_2, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_2, 1); +lean_inc(x_12); +lean_dec(x_2); +x_13 = lean_apply_3(x_5, x_9, x_11, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; +lean_dec(x_5); +lean_dec(x_4); +x_14 = lean_apply_1(x_3, x_2); +return x_14; +} +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___rarg___boxed), 5, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = l___private_Init_Data_List_TakeDrop_0__List_drop_match__1_splitter___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +return x_6; +} +} +lean_object* initialize_Init_Data_List_Lemmas(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_List_TakeDrop(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_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_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/String/Basic.c b/stage0/stdlib/Init/Data/String/Basic.c index 5c350b0c551a..d4a9614b8bc1 100644 --- a/stage0/stdlib/Init/Data/String/Basic.c +++ b/stage0/stdlib/Init/Data/String/Basic.c @@ -77,6 +77,7 @@ LEAN_EXPORT lean_object* l_Substring_front___boxed(lean_object*); LEAN_EXPORT uint8_t l_String_contains(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_length___boxed(lean_object*); LEAN_EXPORT lean_object* l_String_anyAux___at_String_all___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_posOf(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_String_foldl(lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_remainingBytes___boxed(lean_object*); @@ -95,7 +96,6 @@ LEAN_EXPORT lean_object* l_String_firstDiffPos___boxed(lean_object*, lean_object LEAN_EXPORT lean_object* l_String_isNat___boxed(lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Substring_isNat(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_modify___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_String_dropRightWhile(lean_object*, lean_object*); @@ -300,7 +300,6 @@ static lean_object* l_Substring_hasBeq___closed__1; LEAN_EXPORT uint8_t l_String_Iterator_atEnd(lean_object*); LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at_String_nextUntil___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_trimRight(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_toString___boxed(lean_object*); LEAN_EXPORT lean_object* l_String_anyAux___at_Substring_all___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -327,6 +326,7 @@ LEAN_EXPORT lean_object* l_String_foldrAux(lean_object*); LEAN_EXPORT lean_object* l_String_instSizeOfIterator(lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_toEnd(lean_object*); LEAN_EXPORT lean_object* l_String_drop(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_utf8ByteSize_go_match__1_splitter___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_String_Iterator_hasPrev(lean_object*); LEAN_EXPORT lean_object* l_Substring_takeRightWhile(lean_object*, lean_object*); @@ -1868,7 +1868,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531_(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; @@ -1891,11 +1891,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567_(x_1, x_2); +x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1906,7 +1906,7 @@ LEAN_EXPORT uint8_t l_String_instDecidableEqIterator(lean_object* x_1, lean_obje _start: { uint8_t x_3; -x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3567_(x_1, x_2); +x_3 = l___private_Init_Data_String_Basic_0__String_decEqIterator____x40_Init_Data_String_Basic___hyg_3531_(x_1, x_2); return x_3; } } diff --git a/stage0/stdlib/Init/Data/String/Extra.c b/stage0/stdlib/Init/Data/String/Extra.c index 19ed5fe0a55c..cc23abf0d9af 100644 --- a/stage0/stdlib/Init/Data/String/Extra.c +++ b/stage0/stdlib/Init/Data/String/Extra.c @@ -52,6 +52,7 @@ static lean_object* l_String_fromUTF8_x21___closed__1; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__4; LEAN_EXPORT lean_object* l_panic___at_String_toNat_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_String_Iterator_find___at___private_Init_Data_String_Extra_0__String_findLeadingSpacesSize___spec__1(lean_object*); +static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__3; LEAN_EXPORT lean_object* l___private_Init_Data_String_Extra_0__String_findLeadingSpacesSize(lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__1; @@ -101,14 +102,18 @@ lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__6; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15; lean_object* lean_string_to_utf8(lean_object*); +static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__23; +static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25; static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__7; uint32_t lean_uint32_lor(uint32_t, uint32_t); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22; lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24; uint32_t lean_uint32_shift_left(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_validateUTF8_loop(lean_object*, lean_object*); +static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27; LEAN_EXPORT lean_object* l___private_Init_Data_String_Extra_0__List_join_match__1_splitter(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_String_Extra_0__String_findLeadingSpacesSize_consumeSpaces(lean_object*, lean_object*, lean_object*); static lean_object* l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19; @@ -1101,7 +1106,7 @@ static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRu _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("seq1", 4); +x_1 = lean_mk_string_from_bytes("withReducible", 13); return x_1; } } @@ -1121,21 +1126,69 @@ static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRu _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); +x_1 = lean_mk_string_from_bytes("with_reducible", 14); return x_1; } } static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__9() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq", 9); +return x_1; +} +} +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__3; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__4; +x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__5; +x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); +return x_1; +} +} +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__3; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__4; +x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__5; +x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("null", 4); +return x_1; +} +} +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__8; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15() { _start: { lean_object* x_1; @@ -1143,19 +1196,19 @@ x_1 = lean_mk_string_from_bytes("apply", 5); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__11() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__3; x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__4; x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__5; -x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10; +x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__12() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17() { _start: { lean_object* x_1; @@ -1163,16 +1216,16 @@ x_1 = lean_mk_string_from_bytes("String.Iterator.sizeOf_next_lt_of_hasNext", 41) return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__13() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__12; +x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19() { _start: { lean_object* x_1; @@ -1180,7 +1233,7 @@ x_1 = lean_mk_string_from_bytes("String", 6); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20() { _start: { lean_object* x_1; @@ -1188,7 +1241,7 @@ x_1 = lean_mk_string_from_bytes("Iterator", 8); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__16() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21() { _start: { lean_object* x_1; @@ -1196,42 +1249,42 @@ x_1 = lean_mk_string_from_bytes("sizeOf_next_lt_of_hasNext", 25); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14; -x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15; -x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__16; +x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20; +x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__18() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__18; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25() { _start: { lean_object* x_1; @@ -1239,7 +1292,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26() { _start: { lean_object* x_1; @@ -1247,14 +1300,14 @@ x_1 = lean_mk_string_from_bytes("assumption", 10); return x_1; } } -static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22() { +static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__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_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__3; x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__4; x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__5; -x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21; +x_4 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -1277,7 +1330,7 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); x_9 = 0; @@ -1287,46 +1340,57 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); lean_inc(x_12); lean_dec(x_2); -x_13 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10; +x_13 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__8; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__17; -x_16 = l_Lean_addMacroScope(x_12, x_15, x_11); -x_17 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__13; -x_18 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19; +x_15 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15; lean_inc(x_10); -x_19 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_19, 0, x_10); -lean_ctor_set(x_19, 1, x_17); -lean_ctor_set(x_19, 2, x_16); -lean_ctor_set(x_19, 3, x_18); -x_20 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__11; +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22; +x_18 = l_Lean_addMacroScope(x_12, x_17, x_11); +x_19 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__18; +x_20 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24; lean_inc(x_10); -x_21 = l_Lean_Syntax_node2(x_10, x_20, x_14, x_19); -x_22 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20; +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__16; lean_inc(x_10); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_10); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21; +x_23 = l_Lean_Syntax_node2(x_10, x_22, x_16, x_21); +x_24 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25; lean_inc(x_10); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_10); lean_ctor_set(x_25, 1, x_24); -x_26 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22; +x_26 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26; lean_inc(x_10); -x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); -x_28 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__9; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27; lean_inc(x_10); -x_29 = l_Lean_Syntax_node3(x_10, x_28, x_21, x_23, x_27); -x_30 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__7; -x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_3); -return x_32; +x_29 = l_Lean_Syntax_node1(x_10, x_28, x_27); +x_30 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14; +lean_inc(x_10); +x_31 = l_Lean_Syntax_node3(x_10, x_30, x_23, x_25, x_29); +x_32 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__12; +lean_inc(x_10); +x_33 = l_Lean_Syntax_node1(x_10, x_32, x_31); +x_34 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10; +lean_inc(x_10); +x_35 = l_Lean_Syntax_node1(x_10, x_34, x_33); +x_36 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__7; +x_37 = l_Lean_Syntax_node2(x_10, x_36, x_14, 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_3); +return x_38; } } } @@ -1359,8 +1423,8 @@ static lean_object* _init_l_String___aux__Init__Data__String__Extra______macroRu _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14; -x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15; +x_1 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__19; +x_2 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20; x_3 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__3; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; @@ -1408,7 +1472,7 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); x_9 = 0; @@ -1418,46 +1482,57 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); lean_inc(x_12); lean_dec(x_2); -x_13 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10; +x_13 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__8; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__4; -x_16 = l_Lean_addMacroScope(x_12, x_15, x_11); -x_17 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__2; -x_18 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__6; +x_15 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__15; lean_inc(x_10); -x_19 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_19, 0, x_10); -lean_ctor_set(x_19, 1, x_17); -lean_ctor_set(x_19, 2, x_16); -lean_ctor_set(x_19, 3, x_18); -x_20 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__11; +x_16 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_16, 0, x_10); +lean_ctor_set(x_16, 1, x_15); +x_17 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__4; +x_18 = l_Lean_addMacroScope(x_12, x_17, x_11); +x_19 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__2; +x_20 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__6; lean_inc(x_10); -x_21 = l_Lean_Syntax_node2(x_10, x_20, x_14, x_19); -x_22 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__20; +x_21 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_18); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__16; lean_inc(x_10); -x_23 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_23, 0, x_10); -lean_ctor_set(x_23, 1, x_22); -x_24 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21; +x_23 = l_Lean_Syntax_node2(x_10, x_22, x_16, x_21); +x_24 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25; lean_inc(x_10); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_10); lean_ctor_set(x_25, 1, x_24); -x_26 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22; +x_26 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26; lean_inc(x_10); -x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); -x_28 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__9; +x_27 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_27, 0, x_10); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27; lean_inc(x_10); -x_29 = l_Lean_Syntax_node3(x_10, x_28, x_21, x_23, x_27); -x_30 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__7; -x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_3); -return x_32; +x_29 = l_Lean_Syntax_node1(x_10, x_28, x_27); +x_30 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__14; +lean_inc(x_10); +x_31 = l_Lean_Syntax_node3(x_10, x_30, x_23, x_25, x_29); +x_32 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__12; +lean_inc(x_10); +x_33 = l_Lean_Syntax_node1(x_10, x_32, x_31); +x_34 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__10; +lean_inc(x_10); +x_35 = l_Lean_Syntax_node1(x_10, x_34, x_33); +x_36 = l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__7; +x_37 = l_Lean_Syntax_node2(x_10, x_36, x_14, 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_3); +return x_38; } } } @@ -1929,6 +2004,16 @@ l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__tri lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__21); l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22(); lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__22); +l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__23 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__23(); +lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__23); +l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24(); +lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__24); +l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25(); +lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__25); +l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26(); +lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__26); +l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27(); +lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__1___closed__27); l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__1 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__1(); lean_mark_persistent(l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__1); l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__2 = _init_l_String___aux__Init__Data__String__Extra______macroRules__tacticDecreasing__trivial__2___closed__2(); diff --git a/stage0/stdlib/Init/Meta.c b/stage0/stdlib/Init/Meta.c index 947a2187f474..9980ab51a041 100644 --- a/stage0/stdlib/Init/Meta.c +++ b/stage0/stdlib/Init/Meta.c @@ -22,15 +22,15 @@ static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__3; static lean_object* l_Lean_Name_escapePart___closed__2; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEval__prio____1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__26; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17; LEAN_EXPORT lean_object* l_Lean_evalPrio(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_quoteNameMk___closed__7; lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__2; static lean_object* l_Lean_mkGroupNode___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; LEAN_EXPORT lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___boxed(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__4; @@ -41,14 +41,11 @@ static lean_object* l_Lean_Syntax_mkNameLit___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitTerm(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__83; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStr1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_origin___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__26; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__4; @@ -58,28 +55,27 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__63; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__16; LEAN_EXPORT lean_object* l_Lean_instQuoteArrayMkStr1___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__6; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__6; LEAN_EXPORT lean_object* l_Lean_HygieneInfo_mkIdent___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Macro_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___closed__1; uint32_t lean_string_utf8_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getSepElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17443_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push___rarg(lean_object*, lean_object*, lean_object*); @@ -94,7 +90,9 @@ static lean_object* l_Lean_Parser_Tactic_simpArith___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_instReprTSyntax(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__12; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__11; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__5; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__10; LEAN_EXPORT lean_object* l_Lean_TSyntax_getId(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__18; @@ -104,23 +102,19 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_getName(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__84; lean_object* lean_uint32_to_nat(uint32_t); LEAN_EXPORT lean_object* l_Lean_Syntax_instRepr; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__8; LEAN_EXPORT lean_object* l_Lean_evalPrec(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__40; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__52; LEAN_EXPORT lean_object* l_Lean_Name_instRepr; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__25; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__107; static lean_object* l_Lean_Name_reprPrec___closed__5; @@ -129,7 +123,6 @@ LEAN_EXPORT lean_object* lean_name_append_after(lean_object*, lean_object*); static lean_object* l_Lean_TSyntax_getScientific___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__10; static lean_object* l_Lean_termEval__prio_____closed__6; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__15; static lean_object* l_Lean_termEval__prec_____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__39; @@ -138,12 +131,13 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__35; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__3; LEAN_EXPORT lean_object* l_Lean_Name_toString___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getHeadInfo(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_mapSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__16; LEAN_EXPORT lean_object* l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__4; @@ -152,115 +146,112 @@ static lean_object* l_Lean_Parser_Tactic_simpArith___closed__3; static lean_object* l_Lean_Option_hasQuote___rarg___closed__4; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decode(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_location; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2; static lean_object* l_Lean_evalPrec___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionString___closed__7; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNumLit___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__21; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__80; LEAN_EXPORT lean_object* l_Lean_TSyntax_getString(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__37; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__9; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__8; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__113; LEAN_EXPORT lean_object* l_Lean_Syntax_structEq___boxed(lean_object*, lean_object*); static lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray___closed__1; extern lean_object* l_Lean_reservedMacroScope; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22; lean_object* l_String_quote(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__19; LEAN_EXPORT lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13; static lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__134; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__27; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_origin; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__3; lean_object* l_Lean_TSyntaxArray_rawImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_hasQuote(lean_object*); LEAN_EXPORT lean_object* l_Lean_NameGenerator_curr(lean_object*); LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_version_major___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayTSepArray___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__3; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFrom___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Omega_OmegaConfig_splitMinMax___default; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21; static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__1; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10; static lean_object* l_Lean_mkHole___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__46; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__6; LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__91; extern lean_object* l_Lean_Parser_Tactic_simpStar; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12; LEAN_EXPORT lean_object* l_Lean_instQuoteStringStrLitKind(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__32; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__110; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__73; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__79; LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___at_Lean_Syntax_setHeadInfoAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__58; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17; LEAN_EXPORT lean_object* l_Lean_Syntax_isScientificLit_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__12; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm(lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_version_specialDesc___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__19; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7; lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__69; static lean_object* l_Lean_Syntax_instReprTSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); lean_object* l_String_trim(lean_object*); LEAN_EXPORT lean_object* l_Lean_isLetterLike___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__106; static lean_object* l_Lean_termEval__prio_____closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8; lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6; lean_object* l___private_Init_Data_Repr_0__reprSourceInfo____x40_Init_Data_Repr___hyg_2330_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElemsM(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__5; static lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArray___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionSepArray(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__22; uint8_t l_String_contains(lean_object*, uint32_t); @@ -270,25 +261,27 @@ static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__53; LEAN_EXPORT uint8_t l_Lean_isGreek(uint32_t); LEAN_EXPORT uint8_t l_Lean_isIdBeginEscape(uint32_t); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllArith; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__87; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547_(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__5; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_isFieldIdx_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__88; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22; LEAN_EXPORT uint8_t l_Lean_Syntax_instBEqTSyntax___rarg(lean_object*, lean_object*); uint8_t l_Char_isAlpha(uint32_t); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__13; lean_object* lean_version_get_patch(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__75; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20167_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f___boxed(lean_object*); @@ -300,13 +293,14 @@ LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1(lean_objec static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__112; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__36; LEAN_EXPORT uint8_t l_Lean_Meta_Rewrite_Config_transparency___default; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5; static lean_object* l_Lean_Syntax_unsetTrailing___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfoAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isSubScriptAlnum___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_beq___at_Lean_Syntax_structEq___spec__2___boxed(lean_object*, lean_object*); @@ -318,6 +312,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__61; uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6; LEAN_EXPORT uint32_t l_Lean_idBeginEscape; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__25; LEAN_EXPORT lean_object* l_Lean_isGreek___boxed(lean_object*); @@ -326,7 +321,6 @@ static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__114; LEAN_EXPORT lean_object* l_List_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__4(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12; static lean_object* l_Lean_versionString___closed__6; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__3; @@ -334,23 +328,22 @@ LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__6; LEAN_EXPORT lean_object* lean_name_append_before(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__21; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33; uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__1; static lean_object* l_Lean_toolchain___closed__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSepArray___rarg(lean_object*); static lean_object* l_Lean_termEval__prio_____closed__5; static lean_object* l_Lean_termEval__prec_____closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__6; LEAN_EXPORT uint8_t l_Lean_Syntax_isAtom(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__47; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_instReprEtaStructMode; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18; LEAN_EXPORT uint8_t l_Lean_Meta_Rewrite_Config_offsetCnstrs___default; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__3; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrec__1(lean_object*, lean_object*, lean_object*); @@ -358,41 +351,45 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__19; LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2711____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4; LEAN_EXPORT uint8_t l_Lean_version_isRelease; lean_object* lean_string_utf8_byte_size(lean_object*); LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeCharLitTerm___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed__const__1; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__1; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5; lean_object* lean_string_push(lean_object*, uint32_t); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__59; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__2; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__74; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__142; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24; static lean_object* l_Lean_Syntax_instBEq___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16; LEAN_EXPORT lean_object* l_Lean_mkFreshId(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__71; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLitAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295_(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19; LEAN_EXPORT lean_object* l_String_anyAux___at_Lean_Name_escapePart___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_discharger; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__4; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__14; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__4; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3; LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionTSepArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterDot___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__5; LEAN_EXPORT lean_object* l_Lean_Name_toStringWithSep_maybeEscape(uint8_t, lean_object*); @@ -404,7 +401,6 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_splitNameLit(lean_object*); static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__33; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f___boxed(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); @@ -424,11 +420,11 @@ static lean_object* l_Lean_versionStringCore___closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_expandInterpolatedStr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__4; static lean_object* l_Lean_Syntax_mkCharLit___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__9; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__34; size_t lean_usize_of_nat(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; static lean_object* l_Lean_mkOptionalNode___closed__2; uint8_t l_String_isEmpty(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__16; @@ -440,24 +436,23 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__45; LEAN_EXPORT lean_object* l_Lean_instQuoteBoolMkStr1(uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__27; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_instReprConfig; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__96; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__7; static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6; LEAN_EXPORT lean_object* l_Lean_toolchain; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__30; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__37; uint8_t l_instDecidableNot___rarg(uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__124; @@ -466,7 +461,12 @@ static lean_object* l_Lean_Syntax_isFieldIdx_x3f___closed__2; LEAN_EXPORT uint8_t l_Lean_Meta_Omega_OmegaConfig_splitNatSub___default; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKind___rarg(lean_object*); static lean_object* l_Array_getSepElems___rarg___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10; static lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3; static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__3; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__21; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__25; @@ -475,18 +475,19 @@ static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Synt lean_object* lean_string_utf8_next(lean_object*, lean_object*); static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__2; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__14; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__89; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__105; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__15; uint8_t l_Char_isAlphanum(uint32_t); LEAN_EXPORT lean_object* l_Lean_Meta_instReprTransparencyMode; uint8_t l_Char_isWhitespace(uint32_t); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElems___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32; lean_object* l_String_dropRight(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19217_(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__126; static lean_object* l_Lean_version_patch___closed__1; @@ -495,8 +496,6 @@ static lean_object* l_Lean_instQuoteBoolMkStr1___closed__2; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__73; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295_(uint8_t, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__5; @@ -511,6 +510,7 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast(lean_ob lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__17; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20; LEAN_EXPORT uint8_t l_List_beq___at_Lean_Syntax_structEq___spec__2(lean_object*, lean_object*); uint8_t lean_uint32_dec_le(uint32_t, uint32_t); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__3; @@ -522,6 +522,7 @@ static lean_object* l_Lean_quoteNameMk___closed__6; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeStrLitTerm___boxed(lean_object*); LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__64; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5; static lean_object* l_Lean_Syntax_instCoeOutTSepArrayTSyntaxArray___closed__1; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isIdOrAtom_x3f___boxed(lean_object*); @@ -532,7 +533,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_TSyntax_expandInterpolatedStrChunks___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil(lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_reprPrec___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3; LEAN_EXPORT lean_object* l_Lean_mkHole(lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__102; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeBinLitAux___boxed(lean_object*, lean_object*, lean_object*); @@ -540,39 +540,38 @@ static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_repr static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__13; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943_(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073_(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeHexDigit___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__20; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__95; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instCoeIdentTSyntaxConsSyntaxNodeKindMkStr4Nil___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4; LEAN_EXPORT uint32_t l_Lean_TSyntax_getChar(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__85; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__130; LEAN_EXPORT lean_object* l_Lean_versionString; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27; extern lean_object* l_Lean_Parser_Tactic_rwRuleSeq; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__31; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1; static lean_object* l_Lean_termEval__prio_____closed__9; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___lambda__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; extern lean_object* l_Lean_Parser_Tactic_simpLemma; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeAfterExp(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__2; @@ -581,12 +580,14 @@ static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__2; LEAN_EXPORT lean_object* l_Lean_mkCIdentFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__132; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__26; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static uint8_t l_Lean_versionString___closed__2; static lean_object* l_Lean_Name_reprPrec___closed__9; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__65; static lean_object* l_Lean_Syntax_getHead_x3f___closed__3; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__2; @@ -598,25 +599,25 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeDecimalLitAu static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__9; static lean_object* l_Lean_termEval__prec_____closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__56; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__3; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__9; LEAN_EXPORT lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkGroupNode(lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Syntax_structEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__88; LEAN_EXPORT lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeDecimalLitAux___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__131; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4; static lean_object* l_Lean_Syntax_instBEqPreresolved___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCharLit___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1___rarg(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Name_appendIndexAfter___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__70; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__1; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378_(uint8_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__69; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__4; uint8_t l_List_isEmpty___rarg(lean_object*); @@ -633,37 +634,33 @@ static lean_object* l_Lean_versionStringCore___closed__4; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__5; static lean_object* l_Lean_TSyntax_expandInterpolatedStr___closed__3; lean_object* lean_array_to_list(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8; LEAN_EXPORT lean_object* l_Lean_getGithash___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__129; lean_object* l_Lean_TSyntaxArray_mkImpl___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Std_Format_joinSep___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7; LEAN_EXPORT lean_object* l_Substring_takeWhileAux___at___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNumLit___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__22; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__4; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2; static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__9; static lean_object* l_Lean_termEval__prio_____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__19; LEAN_EXPORT uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeDepTermMkIdentIdent(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7; LEAN_EXPORT lean_object* l_Array_filterSepElemsM___at_Array_filterSepElems___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__3; static lean_object* l_Lean_Option_hasQuote___rarg___closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__1; @@ -679,26 +676,24 @@ static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_setTailInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__27; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Substring_toName(lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__127; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; static lean_object* l_Lean_mkOptionalNode___closed__1; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__12; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateLast___at_Lean_Syntax_setTailInfoAux___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__18; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__5; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_instReprConfig__1; static lean_object* l_Lean_Meta_instReprConfig__1___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_beq_match__1_splitter(lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteCharCharLitKind___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_mkLit(lean_object*, lean_object*, lean_object*); @@ -714,7 +709,6 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArray(lean_object*); static lean_object* l_Lean_Name_reprPrec___closed__2; static lean_object* l_Lean_versionString___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__5; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Rewrite_Config_occs___default; LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isIdEndEscape___boxed(lean_object*); @@ -726,12 +720,10 @@ static lean_object* l_Lean_version_minor___closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isStrLit_x3f___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__31; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_findAux___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__9; lean_object* l_Lean_extractMacroScopes(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__4; @@ -740,23 +732,29 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrio(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__57; LEAN_EXPORT lean_object* l_Lean_Syntax_getHead_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__20; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__8; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push___boxed(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__109; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__2; static lean_object* l_Lean_quoteNameMk___closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_copyHeadTailInfoFrom___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkCharLit___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__139; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6; LEAN_EXPORT lean_object* l_Lean_Internal_isStage0___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21; uint8_t lean_version_get_is_release(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_isNumericSubscript___boxed(lean_object*); @@ -769,11 +767,12 @@ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_isIdRest___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f___boxed(lean_object*); static lean_object* l_Lean_TSyntax_getScientific___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__80; LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__subPrio__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23; LEAN_EXPORT lean_object* l_Lean_instQuoteOfCoeHTCTTSyntaxConsSyntaxNodeKindNil___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_version_minor; lean_object* lean_string_utf8_prev(lean_object*, lean_object*); @@ -781,43 +780,48 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkOptionalNode(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__13; static lean_object* l_Lean_toolchain___closed__8; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21; 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___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24; static lean_object* l_Lean_mkSepArray___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__123; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_isIdBeginEscape___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__19; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKindNil___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__61; LEAN_EXPORT lean_object* l_Array_mapSepElemsM(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__7; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__9; lean_object* lean_version_get_special_desc(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__16; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14; static lean_object* l_Lean_Option_hasQuote___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeScientificLitTerm(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__85; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__52; LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStr1(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5; LEAN_EXPORT lean_object* l_Lean_instQuoteNameMkStr1(lean_object*); LEAN_EXPORT lean_object* lean_mk_syntax_ident(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__41; @@ -825,6 +829,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__22; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__46; LEAN_EXPORT lean_object* l_Lean_Name_capitalize(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1; static lean_object* l_Lean_versionString___closed__5; LEAN_EXPORT lean_object* l_Lean_TSyntax_getDocString(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__141; @@ -839,12 +844,10 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_isToken___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterSepElems(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMinor___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35; LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__18; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557_(lean_object*, lean_object*); lean_object* l_Lean_MacroScopesView_review(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeQuotedChar___boxed__const__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -852,17 +855,17 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_isNone___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__62; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22; LEAN_EXPORT lean_object* l_Lean_Name_isInaccessibleUserName___boxed(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14; lean_object* lean_get_githash(lean_object*); static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_isAtom___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__9; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -880,19 +883,19 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f(lean_object*, uint8_t, u LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentTerm___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__3; static lean_object* l_Lean_mkHole___closed__2; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19281_(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__47; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__28; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__25; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__29; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__9; static lean_object* l_Lean_versionString___closed__9; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32; LEAN_EXPORT lean_object* l_Array_mapSepElems(lean_object*, lean_object*); extern lean_object* l_Lean_Parser_Tactic_simpErase; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__43; lean_object* l_Substring_takeWhileAux___at_Substring_trimLeft___spec__1(lean_object*, lean_object*, lean_object*); uint8_t l_Char_isDigit(uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__83; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCharLit(uint32_t, lean_object*); @@ -902,13 +905,11 @@ LEAN_EXPORT lean_object* l_Lean_evalOptPrio(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTSyntaxArrayOfTSyntax___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isLetterLike(uint32_t); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__10; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__2___boxed(lean_object*); lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Substring_nextn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* lean_name_append_index_after(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9; LEAN_EXPORT uint8_t l_String_anyAux___at_Lean_Name_escapePart___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_decodeStrLit___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__29; @@ -916,6 +917,7 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Array_mapSepElemsM___at_Array_mapSepElems___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__76; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__24; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14; static lean_object* l_Lean_instQuoteProdMkStr1___rarg___closed__2; static lean_object* l_Lean_quoteNameMk___closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_getElems___rarg___boxed(lean_object*); @@ -923,15 +925,14 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Name_toString_maybePseudoSyntax___closed__1; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__7; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14; LEAN_EXPORT lean_object* l_Lean_TSyntax_getId___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNameLit_x3f(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__13; static lean_object* l_Lean_Syntax_getHead_x3f___closed__2; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____boxed(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; static lean_object* l_Lean_Name_reprPrec___closed__10; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8; static lean_object* l_Lean_toolchain___closed__9; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__74; @@ -941,6 +942,7 @@ LEAN_EXPORT uint8_t l___private_Init_Meta_0__Lean_Name_hasNum(lean_object*); static lean_object* l_Lean_Syntax_mkStrLit___closed__2; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__26; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_updateFirst___at_Lean_Syntax_setHeadInfoAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; static lean_object* l_Lean_quoteNameMk___closed__2; LEAN_EXPORT uint8_t l_Lean_Syntax_isToken(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__6; @@ -948,7 +950,9 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__1; static lean_object* l_Lean_Name_reprPrec___closed__6; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__136; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__7; static lean_object* l_Lean_toolchain___closed__2; LEAN_EXPORT lean_object* l_Lean_TSyntax_getDocString___boxed(lean_object*); @@ -960,7 +964,6 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean uint8_t l_Lean_Name_hasMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f(lean_object*); LEAN_EXPORT uint32_t l_Lean_idEndEscape; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Syntax_instReprPreresolved___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Occurrences_contains___boxed(lean_object*, lean_object*); @@ -975,38 +978,41 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_Name_reprPrec___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__16; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeOctalLitAux___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_repr___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__4___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21; LEAN_EXPORT lean_object* l_repr___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__2(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__66; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__15; LEAN_EXPORT uint8_t l_Lean_Name_toString_maybePseudoSyntax(lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__12; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10; static lean_object* l_Lean_instQuoteSubstringMkStr1___closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__24; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); uint8_t lean_internal_is_stage0(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__28; LEAN_EXPORT lean_object* l_Lean_Syntax_getSubstring_x3f___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__30; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getScientific(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__4; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__20; lean_object* lean_string_length(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5; LEAN_EXPORT lean_object* l_Lean_termEval__prio__; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Name_hasNum___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_getTailInfo_x3f___boxed(lean_object*); static lean_object* l_Lean_Name_toStringWithSep___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_getOptionalIdent_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_mapSepElems___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__30; static lean_object* l_Lean_Name_isInaccessibleUserName___closed__1; LEAN_EXPORT uint8_t l_Lean_Syntax_hasArgs(lean_object*); @@ -1016,7 +1022,7 @@ uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instQuoteOfCoeHTCTTSyntaxConsSyntaxNodeKindNil(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__4; static lean_object* l_List_foldr___at_Substring_toName___spec__1___closed__3; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutSepArrayArray___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__54; @@ -1034,25 +1040,22 @@ lean_object* lean_nat_mod(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__42; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__6; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__20; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__14; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__4; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2; LEAN_EXPORT lean_object* l_Lean_Syntax_getTrailingSize___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__11; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Syntax_find_x3f(lean_object*, lean_object*); uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__119; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8; LEAN_EXPORT lean_object* l_Lean_Internal_hasLLVMBackend___boxed(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7; lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__15; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkIdentFromRef___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__20; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__120; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Syntax_getTailInfo_x3f___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__79; @@ -1061,19 +1064,19 @@ static lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__40; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_getHead_x3f___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__2; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAllKind; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__94; +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_toolchain___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_getElems(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__38; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__6; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__5; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__67; LEAN_EXPORT lean_object* l_Lean_version_patch; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrQuotedChar___boxed(lean_object*, lean_object*); @@ -1091,32 +1094,27 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_findAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKindNil___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); uint8_t l_String_isPrefixOf(lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__21; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__77; LEAN_EXPORT lean_object* l_Lean_Syntax_isCharLit_x3f(lean_object*); lean_object* lean_array_set(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__76; LEAN_EXPORT lean_object* l_Lean_Name_escapePart(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__13; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2; LEAN_EXPORT lean_object* l_Lean_instQuoteSubstringMkStr1___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__86; LEAN_EXPORT lean_object* l_Lean_Syntax_isInterpolatedStrLit_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instBEqTSyntax___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17507_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__14; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_String_toName(lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__21; static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__3; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__108; static lean_object* l_Lean_quoteNameMk___closed__1; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__10; @@ -1126,17 +1124,21 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_decodeScientificLitVal_x3f_decodeExp(lean static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__32; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__10; LEAN_EXPORT lean_object* l_Lean_Name_getRoot___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__8; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailArraySyntaxTSyntaxArray(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__118; LEAN_EXPORT lean_object* l_Lean_Name_eraseSuffix_x3f(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__44; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5; LEAN_EXPORT uint8_t l_Lean_isIdEndEscape(uint32_t); LEAN_EXPORT lean_object* l_Lean_mkFreshId___rarg___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__2; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Syntax_decodeInterpStrLit_loop___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12; LEAN_EXPORT lean_object* l_Lean_Syntax_setInfo(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getName___boxed(lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -1144,18 +1146,16 @@ LEAN_EXPORT uint8_t l_Lean_Name_instDecidableEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_setHeadInfoAux(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkNameLit___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_push(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__66; LEAN_EXPORT lean_object* l_Lean_withHeadRefOnly(lean_object*); static lean_object* l_Lean_mkSepArray___closed__1; static lean_object* l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__7; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794_(lean_object*, lean_object*); static lean_object* l_Lean_mkCIdentFrom___closed__2; lean_object* l_Array_mkArray1___rarg(lean_object*); LEAN_EXPORT uint8_t l_List_beq___at___private_Init_Meta_0__Lean_Syntax_beqPreresolved____x40_Init_Meta___hyg_2711____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Syntax_mkApp___closed__3; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__59; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16; lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__98; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__75; @@ -1163,16 +1163,17 @@ static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init static lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNumLitPrec___boxed(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__18; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpArith; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Syntax_findAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__6; LEAN_EXPORT lean_object* l_Lean_Name_instDecidableEq___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeScientificLitTerm___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6; static lean_object* l_Lean_termEval__prec_____closed__4; static lean_object* l_Lean_Syntax_mkStrLit___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807_(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__82; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__31; LEAN_EXPORT lean_object* l_Lean_versionStringCore; @@ -1182,6 +1183,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeOutTSyntaxArrayArray___boxed(lean_ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__23; LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__125; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__50; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__22; static lean_object* l_Lean_termEval__prec_____closed__7; @@ -1189,7 +1191,7 @@ LEAN_EXPORT lean_object* l_Lean_Syntax_instEmptyCollectionTSepArray___boxed(lean static lean_object* l_Lean_termEval__prec_____closed__5; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__7; static lean_object* l_Lean_githash___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20103_(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Occurrences_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeNameLitTerm(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_tacticErw____; @@ -1207,7 +1209,6 @@ static lean_object* l_Lean_versionStringCore___closed__7; static lean_object* l_Lean_Syntax_instCoeOutSepArrayArray___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__29; lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__4; static lean_object* l_Lean_instQuoteBoolMkStr1___closed__5; static lean_object* l_Lean_instQuoteNameMkStr1___closed__2; @@ -1230,7 +1231,6 @@ static lean_object* l___private_Init_Meta_0__Lean_quoteArray___rarg___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Omega_OmegaConfig_splitNatAbs___default; static lean_object* l_Lean_Parser_Tactic_tacticErw_______closed__10; LEAN_EXPORT lean_object* l_Lean_instQuoteListMkStr1___rarg(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4; LEAN_EXPORT uint8_t l_Lean_Meta_Occurrences_isAll(lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__7; uint8_t lean_internal_has_llvm_backend(lean_object*); @@ -1246,16 +1246,19 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____spec__7(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__18; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__29; static lean_object* l_Array_forInUnsafe_loop___at_Lean_TSyntax_expandInterpolatedStrChunks___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_TSyntax_getScientific___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_SepArray_ofElemsUsingRef___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__92; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7; LEAN_EXPORT lean_object* l_Lean_Option_hasQuote___rarg(lean_object*, lean_object*); lean_object* l_Lean_Macro_expandMacro_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkCIdentFrom___closed__1; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_getNat___boxed(lean_object*); static lean_object* l_Lean_Meta_instReprTransparencyMode___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__42; @@ -1264,13 +1267,11 @@ LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_quoteList(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__54; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__6; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25; LEAN_EXPORT lean_object* l_Lean_version_getSpecialDesc___boxed(lean_object*); static lean_object* l_Lean_mkOptionalNode___closed__3; static lean_object* l_Lean_termEval__prec_____closed__10; -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137_(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__97; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__53; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__36; @@ -1283,22 +1284,24 @@ extern lean_object* l_Lean_Parser_Tactic_config; LEAN_EXPORT lean_object* l_Lean_mkSepArray___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedSyntax; static lean_object* l_Lean_versionStringCore___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__115; lean_object* lean_version_get_minor(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__22; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____boxed(lean_object*, lean_object*); static lean_object* l___private_Init_Meta_0__Lean_Syntax_reprSyntax____x40_Init_Meta___hyg_2056____closed__5; static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__17; +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4; LEAN_EXPORT lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__24; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__51; LEAN_EXPORT lean_object* l_Lean_Syntax_mkCApp(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__4; LEAN_EXPORT lean_object* l_Lean_instQuoteProdMkStr1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCIdent(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20; static lean_object* l_Lean_quoteNameMk___closed__4; LEAN_EXPORT lean_object* l_List_foldr___at_Substring_toName___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__16; @@ -1307,9 +1310,9 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean LEAN_EXPORT lean_object* l_Lean_Syntax_decodeCharLit___boxed(lean_object*); static lean_object* l_Lean_toolchain___closed__5; lean_object* l_Lean_Syntax_setKind(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9; static lean_object* l_Lean_Parser_Tactic_simpArith___closed__8; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; static lean_object* l_Lean_quoteNameMk___closed__3; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___at_Array_filterSepElems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_TSyntax_getChar___boxed(lean_object*); @@ -1321,7 +1324,6 @@ static lean_object* l_Lean_Name_reprPrec___closed__8; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__26; LEAN_EXPORT lean_object* l_Lean_Syntax_decodeNatLitVal_x3f___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3; static lean_object* l_Lean_termEval__prec_____closed__2; LEAN_EXPORT lean_object* l_List_foldl___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__72; @@ -1329,27 +1331,32 @@ LEAN_EXPORT lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__termEva lean_object* lean_nat_pred(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__128; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__3; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__20; lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__17; lean_object* l_panic___at___private_Init_Prelude_0__Lean_assembleParts___spec__1(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__34; LEAN_EXPORT lean_object* l_Lean_Meta_Occurrences_isAll___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__10; static lean_object* l_Lean_instQuoteTermMkStr1___closed__1; static lean_object* l___private_Init_Meta_0__Lean_quoteList___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapSepElemsM___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5; LEAN_EXPORT lean_object* l_Lean_Syntax_instReprTSyntax___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__17; static lean_object* l___private_Init_Meta_0__Lean_Syntax_splitNameLitAux___closed__1; static lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1___closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__72; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__9; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__31; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5; static lean_object* l_Lean_mkGroupNode___closed__1; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_mkSep___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__49; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__93; @@ -1384,40 +1391,34 @@ static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__21; LEAN_EXPORT lean_object* l_Lean_version_getIsRelease___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStr4Nil(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_TSepArray_ofElems___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__13; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__91; LEAN_EXPORT lean_object* l_Array_getSepElems(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__70; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3; static lean_object* l_Lean_termEval__prio_____closed__3; LEAN_EXPORT lean_object* l_Lean_TSyntax_Compat_instCoeTailSyntax___rarg___boxed(lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__78; LEAN_EXPORT lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__12; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__140; LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_simpAutoUnfold; static lean_object* l_Lean_termEval__prio_____closed__4; -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442_(uint8_t, lean_object*); static lean_object* l_Lean_Option_hasQuote___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_mkHole___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22; static lean_object* l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__135; LEAN_EXPORT lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Name_modifyBase(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__tacticErw______1___closed__11; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_mkSepArray___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isIdFirst(uint32_t); static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrio__1___closed__2; +static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__17; LEAN_EXPORT lean_object* l_Lean_version_specialDesc; LEAN_EXPORT lean_object* l_List_repr_x27___at___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18; LEAN_EXPORT lean_object* l_String_anyAux___at_Lean_Name_escapePart___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__60; lean_object* l_String_toSubstring_x27(lean_object*); @@ -1426,13 +1427,12 @@ static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__23; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__27; static lean_object* l_Lean_Parser_Tactic_simpAllKind___closed__8; static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Syntax__addPrec__1___closed__3; -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; LEAN_EXPORT lean_object* l_Lean_Syntax_instBEq; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17; static lean_object* l_Lean_instQuoteNameMkStr1___closed__1; LEAN_EXPORT lean_object* l_Lean_mkCIdentFromRef(lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_version_get_major(lean_object*); +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8; static lean_object* l_Lean_Parser_Tactic_simpAllArith___closed__5; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__43; LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeIdentLevel___boxed(lean_object*); @@ -1445,8 +1445,9 @@ static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean static lean_object* l_Lean_termEval__prec_____closed__11; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__38; LEAN_EXPORT lean_object* l_Lean_isIdFirst___boxed(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___closed__39; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1; static lean_object* l_Lean_Meta_instReprConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_Syntax_toNat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_dsimpAutoUnfold; @@ -1456,7 +1457,9 @@ static lean_object* l_Lean_Syntax_instRepr___closed__1; LEAN_EXPORT lean_object* l_Lean_monadNameGeneratorLift(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_dsimpKind___closed__8; static lean_object* l_Lean_toolchain___closed__1; +static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5; lean_object* l_Char_ofNat(lean_object*); +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__1; static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__133; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getPatch___boxed(lean_object*); @@ -1465,7 +1468,6 @@ LEAN_EXPORT lean_object* l_Lean_TSyntax_instCoeConsSyntaxNodeKind___boxed(lean_o static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__57; static lean_object* l_Lean_Parser_Tactic_simpAutoUnfold___closed__11; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Array_getSepElems___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_declareSimpLikeTactic___closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); static lean_object* l_Lean_Parser_Tactic_simpArith___closed__10; @@ -1477,13 +1479,11 @@ static lean_object* l_Lean___aux__Init__Meta______macroRules__Lean__Parser__Synt static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__45; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Array_filterSepElemsMAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_versionStringCore___closed__6; -static lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12; LEAN_EXPORT lean_object* l_Lean_TSyntax_getHygieneInfo___boxed(lean_object*); static lean_object* l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tactic__declareSimpLikeTactic__1___lambda__1___closed__14; LEAN_EXPORT lean_object* l_Lean_termEval__prec__; static lean_object* l_Lean_Syntax_mkApp___closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); -static lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7; LEAN_EXPORT lean_object* l_Lean_Syntax_instCoeTermTSyntaxConsSyntaxNodeKindMkStr4Nil___boxed(lean_object*); static lean_object* l_Lean_versionStringCore___closed__8; LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_version_getMajor___boxed(lean_object* x_1) { @@ -18692,7 +18692,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1() { _start: { lean_object* x_1; @@ -18700,33 +18700,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.all", 30); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18734,23 +18734,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18758,7 +18758,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7() { _start: { lean_object* x_1; @@ -18766,33 +18766,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.default", 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18800,23 +18800,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18824,7 +18824,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13() { _start: { lean_object* x_1; @@ -18832,33 +18832,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.reducible", 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18866,23 +18866,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18890,7 +18890,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19() { _start: { lean_object* x_1; @@ -18898,33 +18898,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.TransparencyMode.instances", 36); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18932,23 +18932,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -18956,7 +18956,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -18968,14 +18968,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -18988,14 +18988,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -19008,14 +19008,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -19028,14 +19028,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24; +x_25 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -19043,13 +19043,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -19058,7 +19058,7 @@ static lean_object* _init_l_Lean_Meta_instReprTransparencyMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____boxed), 2, 0); return x_1; } } @@ -19070,7 +19070,7 @@ x_1 = l_Lean_Meta_instReprTransparencyMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1() { _start: { lean_object* x_1; @@ -19078,33 +19078,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.all", 27); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19112,23 +19112,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19136,7 +19136,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7() { _start: { lean_object* x_1; @@ -19144,33 +19144,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.notClasses", 34); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19178,23 +19178,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19202,7 +19202,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13() { _start: { lean_object* x_1; @@ -19210,33 +19210,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.EtaStructMode.none", 28); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__4; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19244,23 +19244,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Init_Meta_0__Lean_Syntax_reprPreresolved____x40_Init_Meta___hyg_1943____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19268,7 +19268,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -19280,14 +19280,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4; +x_5 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6; +x_7 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -19300,14 +19300,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10; +x_11 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12; +x_13 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -19320,14 +19320,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16; +x_17 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18; +x_19 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -19335,13 +19335,13 @@ return x_20; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442_(x_3, x_2); +x_4 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -19350,7 +19350,7 @@ static lean_object* _init_l_Lean_Meta_instReprEtaStructMode___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____boxed), 2, 0); return x_1; } } @@ -19362,7 +19362,7 @@ x_1 = l_Lean_Meta_instReprEtaStructMode___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1() { _start: { lean_object* x_1; @@ -19370,33 +19370,33 @@ x_1 = lean_mk_string_from_bytes("zeta", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19404,7 +19404,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -19413,7 +19413,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6() { _start: { lean_object* x_1; @@ -19421,17 +19421,17 @@ x_1 = lean_mk_string_from_bytes("beta", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8() { _start: { lean_object* x_1; @@ -19439,17 +19439,17 @@ x_1 = lean_mk_string_from_bytes("eta", 3); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10() { _start: { lean_object* x_1; @@ -19457,17 +19457,17 @@ x_1 = lean_mk_string_from_bytes("etaStruct", 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -19476,7 +19476,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13() { _start: { lean_object* x_1; @@ -19484,17 +19484,17 @@ x_1 = lean_mk_string_from_bytes("iota", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15() { _start: { lean_object* x_1; @@ -19502,17 +19502,17 @@ x_1 = lean_mk_string_from_bytes("proj", 4); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17() { _start: { lean_object* x_1; @@ -19520,17 +19520,17 @@ x_1 = lean_mk_string_from_bytes("decide", 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -19539,7 +19539,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20() { _start: { lean_object* x_1; @@ -19547,17 +19547,17 @@ x_1 = lean_mk_string_from_bytes("autoUnfold", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22() { _start: { lean_object* x_1; lean_object* x_2; @@ -19566,7 +19566,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23() { _start: { lean_object* x_1; @@ -19574,17 +19574,17 @@ x_1 = lean_mk_string_from_bytes("failIfUnchanged", 15); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25() { _start: { lean_object* x_1; lean_object* x_2; @@ -19593,7 +19593,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26() { _start: { lean_object* x_1; @@ -19601,17 +19601,17 @@ x_1 = lean_mk_string_from_bytes("unfoldPartialApp", 16); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28() { _start: { lean_object* x_1; lean_object* x_2; @@ -19620,7 +19620,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29() { _start: { lean_object* x_1; @@ -19628,17 +19628,17 @@ x_1 = lean_mk_string_from_bytes("zetaDelta", 9); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31() { _start: { lean_object* x_1; lean_object* x_2; @@ -19648,23 +19648,23 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19672,7 +19672,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34() { _start: { lean_object* x_1; lean_object* x_2; @@ -19682,23 +19682,23 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -19706,7 +19706,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; uint8_t x_4; uint8_t x_5; uint8_t x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; uint8_t x_15; uint8_t x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; @@ -19715,8 +19715,8 @@ x_4 = lean_ctor_get_uint8(x_1, 1); x_5 = lean_ctor_get_uint8(x_1, 2); x_6 = lean_ctor_get_uint8(x_1, 3); x_7 = lean_unsigned_to_nat(0u); -x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442_(x_6, x_7); -x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12; +x_8 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378_(x_6, x_7); +x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12; x_10 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -19734,28 +19734,28 @@ x_19 = lean_ctor_get_uint8(x_1, 10); if (x_3 == 0) { lean_object* x_160; -x_160 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_160 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_20 = x_160; goto block_159; } else { lean_object* x_161; -x_161 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_161 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_20 = x_161; goto block_159; } block_159: { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_21 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5; +x_21 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5; x_22 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_23, 0, x_22); lean_ctor_set_uint8(x_23, sizeof(void*)*1, x_11); -x_24 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4; +x_24 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4; x_25 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); @@ -19767,7 +19767,7 @@ x_28 = lean_box(1); x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7; +x_30 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7; x_31 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); @@ -19778,14 +19778,14 @@ lean_ctor_set(x_33, 1, x_32); if (x_4 == 0) { lean_object* x_157; -x_157 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_157 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_34 = x_157; goto block_156; } else { lean_object* x_158; -x_158 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_158 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_34 = x_158; goto block_156; } @@ -19807,7 +19807,7 @@ lean_ctor_set(x_38, 1, x_26); x_39 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_39, 0, x_38); lean_ctor_set(x_39, 1, x_28); -x_40 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9; +x_40 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9; x_41 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_41, 0, x_39); lean_ctor_set(x_41, 1, x_40); @@ -19817,14 +19817,14 @@ lean_ctor_set(x_42, 1, x_32); if (x_5 == 0) { lean_object* x_154; -x_154 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_154 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_43 = x_154; goto block_153; } else { lean_object* x_155; -x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_155 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_43 = x_155; goto block_153; } @@ -19847,7 +19847,7 @@ lean_ctor_set(x_48, 1, x_26); x_49 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_49, 0, x_48); lean_ctor_set(x_49, 1, x_28); -x_50 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11; +x_50 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11; x_51 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); @@ -19863,7 +19863,7 @@ lean_ctor_set(x_54, 1, x_26); x_55 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_55, 0, x_54); lean_ctor_set(x_55, 1, x_28); -x_56 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14; +x_56 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14; x_57 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_57, 0, x_55); lean_ctor_set(x_57, 1, x_56); @@ -19873,14 +19873,14 @@ lean_ctor_set(x_58, 1, x_32); if (x_13 == 0) { lean_object* x_151; -x_151 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_151 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_59 = x_151; goto block_150; } else { lean_object* x_152; -x_152 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_152 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_59 = x_152; goto block_150; } @@ -19902,7 +19902,7 @@ lean_ctor_set(x_63, 1, x_26); x_64 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_64, 0, x_63); lean_ctor_set(x_64, 1, x_28); -x_65 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16; +x_65 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16; x_66 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); @@ -19912,14 +19912,14 @@ lean_ctor_set(x_67, 1, x_32); if (x_14 == 0) { lean_object* x_148; -x_148 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_148 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_68 = x_148; goto block_147; } else { lean_object* x_149; -x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_149 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_68 = x_149; goto block_147; } @@ -19941,7 +19941,7 @@ lean_ctor_set(x_72, 1, x_26); x_73 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_73, 0, x_72); lean_ctor_set(x_73, 1, x_28); -x_74 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18; +x_74 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18; x_75 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_75, 0, x_73); lean_ctor_set(x_75, 1, x_74); @@ -19951,21 +19951,21 @@ lean_ctor_set(x_76, 1, x_32); if (x_15 == 0) { lean_object* x_145; -x_145 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_145 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_77 = x_145; goto block_144; } else { lean_object* x_146; -x_146 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_146 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_77 = x_146; goto block_144; } block_144: { 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_78 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19; +x_78 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19; x_79 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_79, 0, x_78); lean_ctor_set(x_79, 1, x_77); @@ -19981,7 +19981,7 @@ lean_ctor_set(x_82, 1, x_26); x_83 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_83, 0, x_82); lean_ctor_set(x_83, 1, x_28); -x_84 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21; +x_84 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21; x_85 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); @@ -19991,21 +19991,21 @@ lean_ctor_set(x_86, 1, x_32); if (x_16 == 0) { lean_object* x_142; -x_142 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_142 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_87 = x_142; goto block_141; } else { lean_object* x_143; -x_143 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_143 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_87 = x_143; goto block_141; } block_141: { 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_88 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22; +x_88 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22; x_89 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_87); @@ -20021,7 +20021,7 @@ lean_ctor_set(x_92, 1, x_26); x_93 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_93, 0, x_92); lean_ctor_set(x_93, 1, x_28); -x_94 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24; +x_94 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24; x_95 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_95, 0, x_93); lean_ctor_set(x_95, 1, x_94); @@ -20031,21 +20031,21 @@ lean_ctor_set(x_96, 1, x_32); if (x_17 == 0) { lean_object* x_139; -x_139 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_139 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_97 = x_139; goto block_138; } else { lean_object* x_140; -x_140 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_140 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_97 = x_140; goto block_138; } block_138: { lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_98 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25; +x_98 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25; x_99 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_99, 0, x_98); lean_ctor_set(x_99, 1, x_97); @@ -20061,7 +20061,7 @@ lean_ctor_set(x_102, 1, x_26); x_103 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_103, 0, x_102); lean_ctor_set(x_103, 1, x_28); -x_104 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27; +x_104 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27; x_105 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_105, 0, x_103); lean_ctor_set(x_105, 1, x_104); @@ -20071,21 +20071,21 @@ lean_ctor_set(x_106, 1, x_32); if (x_18 == 0) { lean_object* x_136; -x_136 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_136 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_107 = x_136; goto block_135; } else { lean_object* x_137; -x_137 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_137 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_107 = x_137; goto block_135; } block_135: { lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_108 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28; +x_108 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28; x_109 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_109, 0, x_108); lean_ctor_set(x_109, 1, x_107); @@ -20101,7 +20101,7 @@ lean_ctor_set(x_112, 1, x_26); x_113 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_113, 0, x_112); lean_ctor_set(x_113, 1, x_28); -x_114 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30; +x_114 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30; x_115 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); @@ -20111,7 +20111,7 @@ lean_ctor_set(x_116, 1, x_32); if (x_19 == 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; -x_117 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33; +x_117 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33; x_118 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_118, 0, x_116); lean_ctor_set(x_118, 1, x_117); @@ -20135,7 +20135,7 @@ return x_125; else { lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_126 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36; +x_126 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36; x_127 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_127, 0, x_116); lean_ctor_set(x_127, 1, x_126); @@ -20167,11 +20167,11 @@ return x_134; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -20181,7 +20181,7 @@ static lean_object* _init_l_Lean_Meta_instReprConfig___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____boxed), 2, 0); return x_1; } } @@ -20193,7 +20193,7 @@ x_1 = l_Lean_Meta_instReprConfig___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1() { _start: { lean_object* x_1; @@ -20201,33 +20201,33 @@ x_1 = lean_mk_string_from_bytes("maxSteps", 8); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3; x_2 = l___private_Init_Meta_0__Lean_Syntax_reprTSyntax____x40_Init_Meta___hyg_2295____rarg___closed__5; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20235,7 +20235,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -20244,7 +20244,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6() { _start: { lean_object* x_1; @@ -20252,17 +20252,17 @@ x_1 = lean_mk_string_from_bytes("maxDischargeDepth", 17); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -20271,7 +20271,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9() { _start: { lean_object* x_1; @@ -20279,17 +20279,17 @@ x_1 = lean_mk_string_from_bytes("contextual", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11() { _start: { lean_object* x_1; @@ -20297,17 +20297,17 @@ x_1 = lean_mk_string_from_bytes("memoize", 7); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13() { _start: { lean_object* x_1; lean_object* x_2; @@ -20316,7 +20316,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14() { _start: { lean_object* x_1; @@ -20324,17 +20324,17 @@ x_1 = lean_mk_string_from_bytes("singlePass", 10); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16() { _start: { lean_object* x_1; @@ -20342,17 +20342,17 @@ x_1 = lean_mk_string_from_bytes("arith", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18() { _start: { lean_object* x_1; lean_object* x_2; @@ -20361,7 +20361,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19() { _start: { lean_object* x_1; @@ -20369,17 +20369,17 @@ x_1 = lean_mk_string_from_bytes("dsimp", 5); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21() { _start: { lean_object* x_1; @@ -20387,17 +20387,17 @@ x_1 = lean_mk_string_from_bytes("ground", 6); return x_1; } } -static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22() { +static lean_object* _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21; 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_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; @@ -20406,7 +20406,7 @@ lean_inc(x_3); x_4 = l___private_Init_Data_Repr_0__Nat_reprFast(x_3); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5; +x_6 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -20414,7 +20414,7 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4; +x_10 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -20426,7 +20426,7 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7; +x_16 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -20439,7 +20439,7 @@ lean_inc(x_20); x_21 = l___private_Init_Data_Repr_0__Nat_reprFast(x_20); x_22 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_22, 0, x_21); -x_23 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8; +x_23 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8; x_24 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_24, 0, x_23); lean_ctor_set(x_24, 1, x_22); @@ -20455,7 +20455,7 @@ lean_ctor_set(x_27, 1, x_12); x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_14); -x_29 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10; +x_29 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10; x_30 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_30, 0, x_28); lean_ctor_set(x_30, 1, x_29); @@ -20470,8 +20470,8 @@ x_36 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 4); x_37 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 5); x_38 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 6); x_39 = lean_unsigned_to_nat(0u); -x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442_(x_38, x_39); -x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12; +x_40 = l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378_(x_38, x_39); +x_41 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12; x_42 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_42, 0, x_41); lean_ctor_set(x_42, 1, x_40); @@ -20492,21 +20492,21 @@ lean_dec(x_1); if (x_32 == 0) { lean_object* x_264; -x_264 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_264 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_54 = x_264; goto block_263; } else { lean_object* x_265; -x_265 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_265 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_54 = x_265; goto block_263; } block_263: { lean_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_55 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22; +x_55 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22; x_56 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_56, 0, x_55); lean_ctor_set(x_56, 1, x_54); @@ -20522,7 +20522,7 @@ lean_ctor_set(x_59, 1, x_12); x_60 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_60, 0, x_59); lean_ctor_set(x_60, 1, x_14); -x_61 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12; +x_61 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12; x_62 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_62, 0, x_60); lean_ctor_set(x_62, 1, x_61); @@ -20532,21 +20532,21 @@ lean_ctor_set(x_63, 1, x_18); if (x_33 == 0) { lean_object* x_261; -x_261 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_261 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_64 = x_261; goto block_260; } else { lean_object* x_262; -x_262 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_262 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_64 = x_262; goto block_260; } block_260: { 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_65 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13; +x_65 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13; x_66 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_66, 0, x_65); lean_ctor_set(x_66, 1, x_64); @@ -20562,7 +20562,7 @@ lean_ctor_set(x_69, 1, x_12); x_70 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_70, 0, x_69); lean_ctor_set(x_70, 1, x_14); -x_71 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15; +x_71 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15; x_72 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_72, 0, x_70); lean_ctor_set(x_72, 1, x_71); @@ -20572,14 +20572,14 @@ lean_ctor_set(x_73, 1, x_18); if (x_34 == 0) { lean_object* x_258; -x_258 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_258 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_74 = x_258; goto block_257; } else { lean_object* x_259; -x_259 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_259 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_74 = x_259; goto block_257; } @@ -20601,7 +20601,7 @@ lean_ctor_set(x_78, 1, x_12); x_79 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_79, 0, x_78); lean_ctor_set(x_79, 1, x_14); -x_80 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2; +x_80 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2; x_81 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_81, 0, x_79); lean_ctor_set(x_81, 1, x_80); @@ -20611,21 +20611,21 @@ lean_ctor_set(x_82, 1, x_18); if (x_35 == 0) { lean_object* x_255; -x_255 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_255 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_83 = x_255; goto block_254; } else { lean_object* x_256; -x_256 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_256 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_83 = x_256; goto block_254; } block_254: { lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_84 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5; +x_84 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5; x_85 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_85, 0, x_84); lean_ctor_set(x_85, 1, x_83); @@ -20641,7 +20641,7 @@ lean_ctor_set(x_88, 1, x_12); x_89 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_89, 0, x_88); lean_ctor_set(x_89, 1, x_14); -x_90 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7; +x_90 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7; x_91 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_91, 0, x_89); lean_ctor_set(x_91, 1, x_90); @@ -20651,14 +20651,14 @@ lean_ctor_set(x_92, 1, x_18); if (x_36 == 0) { lean_object* x_252; -x_252 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_252 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_93 = x_252; goto block_251; } else { lean_object* x_253; -x_253 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_253 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_93 = x_253; goto block_251; } @@ -20680,7 +20680,7 @@ lean_ctor_set(x_97, 1, x_12); x_98 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_98, 0, x_97); lean_ctor_set(x_98, 1, x_14); -x_99 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9; +x_99 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9; x_100 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_100, 0, x_98); lean_ctor_set(x_100, 1, x_99); @@ -20690,14 +20690,14 @@ lean_ctor_set(x_101, 1, x_18); if (x_37 == 0) { lean_object* x_249; -x_249 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_249 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_102 = x_249; goto block_248; } else { lean_object* x_250; -x_250 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_250 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_102 = x_250; goto block_248; } @@ -20720,7 +20720,7 @@ lean_ctor_set(x_107, 1, x_12); x_108 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_108, 0, x_107); lean_ctor_set(x_108, 1, x_14); -x_109 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11; +x_109 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11; x_110 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_110, 0, x_108); lean_ctor_set(x_110, 1, x_109); @@ -20736,7 +20736,7 @@ lean_ctor_set(x_113, 1, x_12); x_114 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_114, 0, x_113); lean_ctor_set(x_114, 1, x_14); -x_115 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14; +x_115 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14; x_116 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_116, 0, x_114); lean_ctor_set(x_116, 1, x_115); @@ -20746,14 +20746,14 @@ lean_ctor_set(x_117, 1, x_18); if (x_44 == 0) { lean_object* x_246; -x_246 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_246 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_118 = x_246; goto block_245; } else { lean_object* x_247; -x_247 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_247 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_118 = x_247; goto block_245; } @@ -20775,7 +20775,7 @@ lean_ctor_set(x_122, 1, x_12); x_123 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_123, 0, x_122); lean_ctor_set(x_123, 1, x_14); -x_124 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16; +x_124 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16; x_125 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_125, 0, x_123); lean_ctor_set(x_125, 1, x_124); @@ -20785,14 +20785,14 @@ lean_ctor_set(x_126, 1, x_18); if (x_45 == 0) { lean_object* x_243; -x_243 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_243 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_127 = x_243; goto block_242; } else { lean_object* x_244; -x_244 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_244 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_127 = x_244; goto block_242; } @@ -20814,7 +20814,7 @@ lean_ctor_set(x_131, 1, x_12); x_132 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_132, 0, x_131); lean_ctor_set(x_132, 1, x_14); -x_133 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18; +x_133 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18; x_134 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_134, 0, x_132); lean_ctor_set(x_134, 1, x_133); @@ -20824,21 +20824,21 @@ lean_ctor_set(x_135, 1, x_18); if (x_46 == 0) { lean_object* x_240; -x_240 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_240 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_136 = x_240; goto block_239; } else { lean_object* x_241; -x_241 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_241 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_136 = x_241; goto block_239; } block_239: { lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_137 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19; +x_137 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19; x_138 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_138, 0, x_137); lean_ctor_set(x_138, 1, x_136); @@ -20854,7 +20854,7 @@ lean_ctor_set(x_141, 1, x_12); x_142 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_142, 0, x_141); lean_ctor_set(x_142, 1, x_14); -x_143 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17; +x_143 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17; x_144 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_144, 0, x_142); lean_ctor_set(x_144, 1, x_143); @@ -20864,21 +20864,21 @@ lean_ctor_set(x_145, 1, x_18); if (x_47 == 0) { lean_object* x_237; -x_237 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_237 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_146 = x_237; goto block_236; } else { lean_object* x_238; -x_238 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_238 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_146 = x_238; goto block_236; } block_236: { 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_147 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18; +x_147 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18; x_148 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_148, 0, x_147); lean_ctor_set(x_148, 1, x_146); @@ -20894,7 +20894,7 @@ lean_ctor_set(x_151, 1, x_12); x_152 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_152, 0, x_151); lean_ctor_set(x_152, 1, x_14); -x_153 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21; +x_153 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21; x_154 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_154, 0, x_152); lean_ctor_set(x_154, 1, x_153); @@ -20904,14 +20904,14 @@ lean_ctor_set(x_155, 1, x_18); if (x_48 == 0) { lean_object* x_234; -x_234 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_234 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_156 = x_234; goto block_233; } else { lean_object* x_235; -x_235 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_235 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_156 = x_235; goto block_233; } @@ -20933,7 +20933,7 @@ lean_ctor_set(x_160, 1, x_12); x_161 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_161, 0, x_160); lean_ctor_set(x_161, 1, x_14); -x_162 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20; +x_162 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20; x_163 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_163, 0, x_161); lean_ctor_set(x_163, 1, x_162); @@ -20943,14 +20943,14 @@ lean_ctor_set(x_164, 1, x_18); if (x_49 == 0) { lean_object* x_231; -x_231 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_231 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_165 = x_231; goto block_230; } else { lean_object* x_232; -x_232 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_232 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_165 = x_232; goto block_230; } @@ -20972,7 +20972,7 @@ lean_ctor_set(x_169, 1, x_12); x_170 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_170, 0, x_169); lean_ctor_set(x_170, 1, x_14); -x_171 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24; +x_171 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24; x_172 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_172, 0, x_170); lean_ctor_set(x_172, 1, x_171); @@ -20982,21 +20982,21 @@ lean_ctor_set(x_173, 1, x_18); if (x_50 == 0) { lean_object* x_228; -x_228 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_228 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_174 = x_228; goto block_227; } else { lean_object* x_229; -x_229 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_229 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_174 = x_229; goto block_227; } block_227: { lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_175 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25; +x_175 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25; x_176 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_176, 0, x_175); lean_ctor_set(x_176, 1, x_174); @@ -21012,7 +21012,7 @@ lean_ctor_set(x_179, 1, x_12); x_180 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_180, 0, x_179); lean_ctor_set(x_180, 1, x_14); -x_181 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22; +x_181 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22; x_182 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_182, 0, x_180); lean_ctor_set(x_182, 1, x_181); @@ -21022,14 +21022,14 @@ lean_ctor_set(x_183, 1, x_18); if (x_51 == 0) { lean_object* x_225; -x_225 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_225 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_184 = x_225; goto block_224; } else { lean_object* x_226; -x_226 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_226 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_184 = x_226; goto block_224; } @@ -21051,7 +21051,7 @@ lean_ctor_set(x_188, 1, x_12); x_189 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_189, 0, x_188); lean_ctor_set(x_189, 1, x_14); -x_190 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27; +x_190 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27; x_191 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_191, 0, x_189); lean_ctor_set(x_191, 1, x_190); @@ -21061,21 +21061,21 @@ lean_ctor_set(x_192, 1, x_18); if (x_52 == 0) { lean_object* x_222; -x_222 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31; +x_222 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31; x_193 = x_222; goto block_221; } else { lean_object* x_223; -x_223 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34; +x_223 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34; x_193 = x_223; goto block_221; } block_221: { lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_194 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28; +x_194 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28; x_195 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_195, 0, x_194); lean_ctor_set(x_195, 1, x_193); @@ -21091,7 +21091,7 @@ lean_ctor_set(x_198, 1, x_12); x_199 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_199, 0, x_198); lean_ctor_set(x_199, 1, x_14); -x_200 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30; +x_200 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30; x_201 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_201, 0, x_199); lean_ctor_set(x_201, 1, x_200); @@ -21101,7 +21101,7 @@ lean_ctor_set(x_202, 1, x_18); if (x_53 == 0) { lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; -x_203 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33; +x_203 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33; x_204 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_204, 0, x_202); lean_ctor_set(x_204, 1, x_203); @@ -21125,7 +21125,7 @@ return x_211; else { 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___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36; +x_212 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36; x_213 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_213, 0, x_202); lean_ctor_set(x_213, 1, x_212); @@ -21163,11 +21163,11 @@ return x_220; } } } -LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794_(x_1, x_2); +x_3 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -21176,7 +21176,7 @@ static lean_object* _init_l_Lean_Meta_instReprConfig__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____boxed), 2, 0); return x_1; } } @@ -22212,7 +22212,7 @@ static lean_object* _init_l_Lean_Parser_Tactic_dsimpKind___closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -24700,7 +24700,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_2 = l_String_toSubstring_x27(x_1); return x_2; } @@ -24710,7 +24710,7 @@ static lean_object* _init_l_Lean_Parser_Tactic___aux__Init__Meta______macroRules { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -24722,7 +24722,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; x_2 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__2; x_3 = l_Lean_Parser_Tactic_tacticErw_______closed__1; -x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_4 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -27237,7 +27237,7 @@ x_1 = l_Lean_Parser_Tactic_simpAutoUnfold___closed__26; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -27265,7 +27265,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1() { _start: { lean_object* x_1; @@ -27273,16 +27273,16 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Simp.Config", 21); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3() { _start: { lean_object* x_1; @@ -27290,7 +27290,7 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4() { _start: { lean_object* x_1; @@ -27298,7 +27298,7 @@ x_1 = lean_mk_string_from_bytes("Simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5() { _start: { lean_object* x_1; @@ -27306,84 +27306,84 @@ x_1 = lean_mk_string_from_bytes("Config", 6); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743_(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; @@ -27444,12 +27444,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -27497,11 +27497,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -27568,7 +27568,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -27623,12 +27623,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -27676,11 +27676,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -27732,17 +27732,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -27879,45 +27879,45 @@ x_1 = l_Lean_Parser_Tactic_simpArith___closed__10; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17; +x_1 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5() { _start: { lean_object* x_1; @@ -27925,41 +27925,41 @@ x_1 = lean_mk_string_from_bytes("Decidable", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5; -x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5; +x_2 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547_(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; @@ -28020,12 +28020,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -28073,11 +28073,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -28110,10 +28110,10 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_70 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_70 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_12); x_71 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_71, 0, x_12); @@ -28168,7 +28168,7 @@ x_91 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_92 = l_Lean_Syntax_node2(x_12, x_91, x_84, x_90); x_93 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_92, x_38); -x_94 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_93, x_2, x_3); +x_94 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_93, x_2, x_3); lean_dec(x_2); return x_94; } @@ -28223,12 +28223,12 @@ lean_inc(x_99); x_116 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_116, 0, x_99); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_117 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_100); lean_inc(x_101); x_118 = l_Lean_addMacroScope(x_101, x_117, x_100); -x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_99); x_121 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_121, 0, x_99); @@ -28276,11 +28276,11 @@ lean_ctor_set(x_137, 0, x_99); lean_ctor_set(x_137, 1, x_136); lean_inc(x_99); x_138 = l_Lean_Syntax_node2(x_99, x_122, x_135, x_137); -x_139 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_139 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_100); lean_inc(x_101); x_140 = l_Lean_addMacroScope(x_101, x_139, x_100); -x_141 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_141 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_99); x_142 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_142, 0, x_99); @@ -28313,10 +28313,10 @@ lean_inc(x_99); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_99); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_154 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_155 = l_Lean_addMacroScope(x_101, x_154, x_100); -x_156 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_157 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_156 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_157 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_99); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_99); @@ -28356,7 +28356,7 @@ x_173 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_99); x_174 = l_Lean_Syntax_node2(x_99, x_173, x_171, x_172); x_175 = l_Lean_Syntax_node5(x_99, x_8, x_103, x_105, x_107, x_174, x_125); -x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_175, x_2, x_3); +x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_175, x_2, x_3); lean_dec(x_2); return x_176; } @@ -28494,7 +28494,7 @@ x_1 = l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17507_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_17443_(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; @@ -28555,12 +28555,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -28608,11 +28608,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -28645,11 +28645,11 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_13); lean_inc(x_14); x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_12); x_70 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_70, 0, x_12); @@ -28663,10 +28663,10 @@ lean_inc(x_62); lean_inc(x_20); lean_inc(x_12); x_72 = l_Lean_Syntax_node3(x_12, x_63, x_71, x_20, x_62); -x_73 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_73 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_74 = l_Lean_addMacroScope(x_14, x_73, x_13); -x_75 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_76 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_75 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_76 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_12); x_77 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_77, 0, x_12); @@ -28722,7 +28722,7 @@ x_97 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_98 = l_Lean_Syntax_node2(x_12, x_97, x_90, x_96); x_99 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_98, x_38); -x_100 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_99, x_2, x_3); +x_100 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_99, x_2, x_3); lean_dec(x_2); return x_100; } @@ -28777,12 +28777,12 @@ lean_inc(x_105); x_122 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_122, 0, x_105); lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6; +x_123 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6; lean_inc(x_106); lean_inc(x_107); x_124 = l_Lean_addMacroScope(x_107, x_123, x_106); -x_125 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2; -x_126 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10; +x_125 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2; +x_126 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10; lean_inc(x_105); x_127 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_127, 0, x_105); @@ -28830,11 +28830,11 @@ lean_ctor_set(x_143, 0, x_105); lean_ctor_set(x_143, 1, x_142); lean_inc(x_105); x_144 = l_Lean_Syntax_node2(x_105, x_128, x_141, x_143); -x_145 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_145 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_106); lean_inc(x_107); x_146 = l_Lean_addMacroScope(x_107, x_145, x_106); -x_147 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_147 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_105); x_148 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_148, 0, x_105); @@ -28867,11 +28867,11 @@ lean_inc(x_105); x_159 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_159, 0, x_105); lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_160 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_106); lean_inc(x_107); x_161 = l_Lean_addMacroScope(x_107, x_160, x_106); -x_162 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_162 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_105); x_163 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_163, 0, x_105); @@ -28885,10 +28885,10 @@ lean_inc(x_155); lean_inc(x_113); lean_inc(x_105); x_165 = l_Lean_Syntax_node3(x_105, x_156, x_164, x_113, x_155); -x_166 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_166 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_167 = l_Lean_addMacroScope(x_107, x_166, x_106); -x_168 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_169 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_168 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_169 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_105); x_170 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_170, 0, x_105); @@ -28929,7 +28929,7 @@ x_185 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_105); x_186 = l_Lean_Syntax_node2(x_105, x_185, x_183, x_184); x_187 = l_Lean_Syntax_node5(x_105, x_8, x_109, x_111, x_113, x_186, x_131); -x_188 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____lambda__1(x_1, x_187, x_2, x_3); +x_188 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____lambda__1(x_1, x_187, x_2, x_3); lean_dec(x_2); return x_188; } @@ -29109,7 +29109,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13; return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -29117,7 +29117,7 @@ x_1 = lean_mk_string_from_bytes("simp_all", 8); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -29125,7 +29125,7 @@ x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tact x_6 = l_Lean_Syntax_setKind(x_1, x_5); x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_6, x_7); -x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1; +x_9 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1; x_10 = 1; x_11 = l_Lean_mkAtomFrom(x_8, x_9, x_10); x_12 = l_Lean_Syntax_setArg(x_6, x_7, x_11); @@ -29145,7 +29145,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1() { _start: { lean_object* x_1; @@ -29153,16 +29153,16 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.Simp.ConfigCtx", 24); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3() { _start: { lean_object* x_1; @@ -29170,65 +29170,65 @@ x_1 = lean_mk_string_from_bytes("ConfigCtx", 9); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423_(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; @@ -29289,12 +29289,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -29342,11 +29342,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -29413,7 +29413,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -29468,12 +29468,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -29521,11 +29521,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -29577,17 +29577,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -29710,7 +29710,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArith___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19281_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_19217_(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; @@ -29771,12 +29771,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -29824,11 +29824,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -29861,10 +29861,10 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_70 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_70 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_12); x_71 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_71, 0, x_12); @@ -29919,7 +29919,7 @@ x_91 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_92 = l_Lean_Syntax_node2(x_12, x_91, x_84, x_90); x_93 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_92, x_38); -x_94 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_93, x_2, x_3); +x_94 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_93, x_2, x_3); lean_dec(x_2); return x_94; } @@ -29974,12 +29974,12 @@ lean_inc(x_99); x_116 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_116, 0, x_99); lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_117 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_100); lean_inc(x_101); x_118 = l_Lean_addMacroScope(x_101, x_117, x_100); -x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_119 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_120 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_99); x_121 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_121, 0, x_99); @@ -30027,11 +30027,11 @@ lean_ctor_set(x_137, 0, x_99); lean_ctor_set(x_137, 1, x_136); lean_inc(x_99); x_138 = l_Lean_Syntax_node2(x_99, x_122, x_135, x_137); -x_139 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_139 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_100); lean_inc(x_101); x_140 = l_Lean_addMacroScope(x_101, x_139, x_100); -x_141 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_141 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_99); x_142 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_142, 0, x_99); @@ -30064,10 +30064,10 @@ lean_inc(x_99); x_153 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_153, 0, x_99); lean_ctor_set(x_153, 1, x_152); -x_154 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_154 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_155 = l_Lean_addMacroScope(x_101, x_154, x_100); -x_156 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_157 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_156 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_157 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_99); x_158 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_158, 0, x_99); @@ -30107,7 +30107,7 @@ x_173 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_99); x_174 = l_Lean_Syntax_node2(x_99, x_173, x_171, x_172); x_175 = l_Lean_Syntax_node5(x_99, x_8, x_103, x_105, x_107, x_174, x_125); -x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_175, x_2, x_3); +x_176 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_175, x_2, x_3); lean_dec(x_2); return x_176; } @@ -30231,7 +30231,7 @@ x_1 = l_Lean_Parser_Tactic_simpAllArithAutoUnfold___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20167_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_20103_(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; @@ -30292,12 +30292,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -30345,11 +30345,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -30382,11 +30382,11 @@ lean_inc(x_12); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_12); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_67 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_13); lean_inc(x_14); x_68 = l_Lean_addMacroScope(x_14, x_67, x_13); -x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_69 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_12); x_70 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_70, 0, x_12); @@ -30400,10 +30400,10 @@ lean_inc(x_62); lean_inc(x_20); lean_inc(x_12); x_72 = l_Lean_Syntax_node3(x_12, x_63, x_71, x_20, x_62); -x_73 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_73 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_74 = l_Lean_addMacroScope(x_14, x_73, x_13); -x_75 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_76 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_75 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_76 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_12); x_77 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_77, 0, x_12); @@ -30459,7 +30459,7 @@ x_97 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_98 = l_Lean_Syntax_node2(x_12, x_97, x_90, x_96); x_99 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_98, x_38); -x_100 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_99, x_2, x_3); +x_100 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_99, x_2, x_3); lean_dec(x_2); return x_100; } @@ -30514,12 +30514,12 @@ lean_inc(x_105); x_122 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_122, 0, x_105); lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4; +x_123 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4; lean_inc(x_106); lean_inc(x_107); x_124 = l_Lean_addMacroScope(x_107, x_123, x_106); -x_125 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2; -x_126 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8; +x_125 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2; +x_126 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8; lean_inc(x_105); x_127 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_127, 0, x_105); @@ -30567,11 +30567,11 @@ lean_ctor_set(x_143, 0, x_105); lean_ctor_set(x_143, 1, x_142); lean_inc(x_105); x_144 = l_Lean_Syntax_node2(x_105, x_128, x_141, x_143); -x_145 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2; +x_145 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2; lean_inc(x_106); lean_inc(x_107); x_146 = l_Lean_addMacroScope(x_107, x_145, x_106); -x_147 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1; +x_147 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1; lean_inc(x_105); x_148 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_148, 0, x_105); @@ -30604,11 +30604,11 @@ lean_inc(x_105); x_159 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_159, 0, x_105); lean_ctor_set(x_159, 1, x_158); -x_160 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_160 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_106); lean_inc(x_107); x_161 = l_Lean_addMacroScope(x_107, x_160, x_106); -x_162 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_162 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_105); x_163 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_163, 0, x_105); @@ -30622,10 +30622,10 @@ lean_inc(x_155); lean_inc(x_113); lean_inc(x_105); x_165 = l_Lean_Syntax_node3(x_105, x_156, x_164, x_113, x_155); -x_166 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4; +x_166 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4; x_167 = l_Lean_addMacroScope(x_107, x_166, x_106); -x_168 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3; -x_169 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8; +x_168 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3; +x_169 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8; lean_inc(x_105); x_170 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_170, 0, x_105); @@ -30666,7 +30666,7 @@ x_185 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_105); x_186 = l_Lean_Syntax_node2(x_105, x_185, x_183, x_184); x_187 = l_Lean_Syntax_node5(x_105, x_8, x_109, x_111, x_113, x_186, x_131); -x_188 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1(x_1, x_187, x_2, x_3); +x_188 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1(x_1, x_187, x_2, x_3); lean_dec(x_2); return x_188; } @@ -30804,7 +30804,7 @@ x_1 = l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -30812,7 +30812,7 @@ x_5 = l_Lean_Parser_Tactic___aux__Init__Meta______macroRules__Lean__Parser__Tact x_6 = l_Lean_Syntax_setKind(x_1, x_5); x_7 = lean_unsigned_to_nat(0u); x_8 = l_Lean_Syntax_getArg(x_6, x_7); -x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19; +x_9 = l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19; x_10 = 1; x_11 = l_Lean_mkAtomFrom(x_8, x_9, x_10); x_12 = l_Lean_Syntax_setArg(x_6, x_7, x_11); @@ -30832,7 +30832,7 @@ lean_ctor_set(x_20, 1, x_4); return x_20; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1() { _start: { lean_object* x_1; @@ -30840,16 +30840,16 @@ x_1 = lean_mk_string_from_bytes("Lean.Meta.DSimp.Config", 22); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3() { _start: { lean_object* x_1; @@ -30857,65 +30857,65 @@ x_1 = lean_mk_string_from_bytes("DSimp", 5); return x_1; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Array_mapMUnsafe_map___at_Lean_expandMacros___spec__1___lambda__1___closed__1; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3; -x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3; -x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3; +x_3 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3; +x_4 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8() { +static lean_object* _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5; -x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7; +x_1 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5; +x_2 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137_(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073_(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; @@ -30976,12 +30976,12 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4; +x_30 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4; lean_inc(x_13); lean_inc(x_14); x_31 = l_Lean_addMacroScope(x_14, x_30, x_13); -x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2; -x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8; +x_32 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2; +x_33 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8; lean_inc(x_12); x_34 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_34, 0, x_12); @@ -31029,11 +31029,11 @@ lean_ctor_set(x_50, 0, x_12); lean_ctor_set(x_50, 1, x_49); lean_inc(x_12); x_51 = l_Lean_Syntax_node2(x_12, x_35, x_48, x_50); -x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_52 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_13); lean_inc(x_14); x_53 = l_Lean_addMacroScope(x_14, x_52, x_13); -x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_54 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_12); x_55 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_55, 0, x_12); @@ -31100,7 +31100,7 @@ x_82 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_12); x_83 = l_Lean_Syntax_node2(x_12, x_82, x_75, x_81); x_84 = l_Lean_Syntax_node5(x_12, x_8, x_16, x_18, x_20, x_83, x_38); -x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1(x_1, x_84, x_2, x_3); +x_85 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1(x_1, x_84, x_2, x_3); lean_dec(x_2); return x_85; } @@ -31155,12 +31155,12 @@ lean_inc(x_90); x_107 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_107, 0, x_90); lean_ctor_set(x_107, 1, x_106); -x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4; +x_108 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4; lean_inc(x_91); lean_inc(x_92); x_109 = l_Lean_addMacroScope(x_92, x_108, x_91); -x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2; -x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8; +x_110 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2; +x_111 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8; lean_inc(x_90); x_112 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_112, 0, x_90); @@ -31208,11 +31208,11 @@ lean_ctor_set(x_128, 0, x_90); lean_ctor_set(x_128, 1, x_127); lean_inc(x_90); x_129 = l_Lean_Syntax_node2(x_90, x_113, x_126, x_128); -x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12; +x_130 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12; lean_inc(x_91); lean_inc(x_92); x_131 = l_Lean_addMacroScope(x_92, x_130, x_91); -x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11; +x_132 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11; lean_inc(x_90); x_133 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_133, 0, x_90); @@ -31264,17 +31264,17 @@ x_155 = l_Lean_Syntax_mkApp___closed__2; lean_inc(x_90); x_156 = l_Lean_Syntax_node2(x_90, x_155, x_153, x_154); x_157 = l_Lean_Syntax_node5(x_90, x_8, x_94, x_96, x_98, x_156, x_116); -x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1(x_1, x_157, x_2, x_3); +x_158 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1(x_1, x_157, x_2, x_3); lean_dec(x_2); return x_158; } } } -LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____lambda__1(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____lambda__1(x_1, x_2, x_3, x_4); lean_dec(x_3); return x_5; } @@ -31844,218 +31844,218 @@ l_Lean_TSyntax_expandInterpolatedStr___closed__4 = _init_l_Lean_TSyntax_expandIn lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__4); l_Lean_TSyntax_expandInterpolatedStr___closed__5 = _init_l_Lean_TSyntax_expandInterpolatedStr___closed__5(); lean_mark_persistent(l_Lean_TSyntax_expandInterpolatedStr___closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11295____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprTransparencyMode____x40_Init_Meta___hyg_11231____closed__24); l_Lean_Meta_instReprTransparencyMode___closed__1 = _init_l_Lean_Meta_instReprTransparencyMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprTransparencyMode___closed__1); l_Lean_Meta_instReprTransparencyMode = _init_l_Lean_Meta_instReprTransparencyMode(); lean_mark_persistent(l_Lean_Meta_instReprTransparencyMode); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11442____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprEtaStructMode____x40_Init_Meta___hyg_11378____closed__18); l_Lean_Meta_instReprEtaStructMode___closed__1 = _init_l_Lean_Meta_instReprEtaStructMode___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprEtaStructMode___closed__1); l_Lean_Meta_instReprEtaStructMode = _init_l_Lean_Meta_instReprEtaStructMode(); lean_mark_persistent(l_Lean_Meta_instReprEtaStructMode); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__22); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__23); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__24); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__25); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__26); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__27); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__28); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__29); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__30); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__31); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__32); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__33); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__34); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__35); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11557____closed__36); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__23); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__24); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__25); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__26); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__27); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__28); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__29); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__30); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__31); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__32); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__33); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__34); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__35); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11493____closed__36); l_Lean_Meta_instReprConfig___closed__1 = _init_l_Lean_Meta_instReprConfig___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprConfig___closed__1); l_Lean_Meta_instReprConfig = _init_l_Lean_Meta_instReprConfig(); lean_mark_persistent(l_Lean_Meta_instReprConfig); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__1); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__2); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__3); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__4); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__5); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__6); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__7); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__8); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__9); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__10); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__11); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__12); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__13); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__14); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__15); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__16); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__17); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__18); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__19); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__20); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__21); -l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22(); -lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11794____closed__22); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__1); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__2); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__3); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__4); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__5); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__6); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__7); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__8); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__9); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__10); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__11); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__12); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__13); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__14); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__15); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__16); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__17); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__18); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__19); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__20); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__21); +l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22 = _init_l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22(); +lean_mark_persistent(l___private_Init_Meta_0__Lean_Meta_reprConfig____x40_Init_Meta___hyg_11730____closed__22); l_Lean_Meta_instReprConfig__1___closed__1 = _init_l_Lean_Meta_instReprConfig__1___closed__1(); lean_mark_persistent(l_Lean_Meta_instReprConfig__1___closed__1); l_Lean_Meta_instReprConfig__1 = _init_l_Lean_Meta_instReprConfig__1(); @@ -32806,30 +32806,30 @@ l_Lean_Parser_Tactic_simpAutoUnfold___closed__26 = _init_l_Lean_Parser_Tactic_si lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold___closed__26); l_Lean_Parser_Tactic_simpAutoUnfold = _init_l_Lean_Parser_Tactic_simpAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__8); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__9); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__10); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__11); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15807____closed__12); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__9); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__10); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__11); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_15743____closed__12); l_Lean_Parser_Tactic_simpArith___closed__1 = _init_l_Lean_Parser_Tactic_simpArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__1); l_Lean_Parser_Tactic_simpArith___closed__2 = _init_l_Lean_Parser_Tactic_simpArith___closed__2(); @@ -32852,22 +32852,22 @@ l_Lean_Parser_Tactic_simpArith___closed__10 = _init_l_Lean_Parser_Tactic_simpAri lean_mark_persistent(l_Lean_Parser_Tactic_simpArith___closed__10); l_Lean_Parser_Tactic_simpArith = _init_l_Lean_Parser_Tactic_simpArith(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArith); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16611____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_16547____closed__8); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__1); l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2 = _init_l_Lean_Parser_Tactic_simpArithAutoUnfold___closed__2(); @@ -32918,24 +32918,24 @@ l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13 = _init_l_Lean_Parser_Tactic lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold___closed__13); l_Lean_Parser_Tactic_simpAllAutoUnfold = _init_l_Lean_Parser_Tactic_simpAllAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____lambda__1___closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18487____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____lambda__1___closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_18423____closed__8); l_Lean_Parser_Tactic_simpAllArith___closed__1 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__1(); lean_mark_persistent(l_Lean_Parser_Tactic_simpAllArith___closed__1); l_Lean_Parser_Tactic_simpAllArith___closed__2 = _init_l_Lean_Parser_Tactic_simpAllArith___closed__2(); @@ -32998,22 +32998,22 @@ l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10 = _init_l_Lean_Parser_Tactic_d lean_mark_persistent(l_Lean_Parser_Tactic_dsimpAutoUnfold___closed__10); l_Lean_Parser_Tactic_dsimpAutoUnfold = _init_l_Lean_Parser_Tactic_dsimpAutoUnfold(); lean_mark_persistent(l_Lean_Parser_Tactic_dsimpAutoUnfold); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__1); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__2); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__3); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__4); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__5); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__6); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__7); -l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8(); -lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21137____closed__8); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__1); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__2); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__3); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__4); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__5); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__6); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__7); +l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8 = _init_l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8(); +lean_mark_persistent(l_Lean_Parser_Tactic_expandSimp____x40_Init_Meta___hyg_21073____closed__8); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Notation.c b/stage0/stdlib/Init/Notation.c index ed974ff96f80..7800866ba1ae 100644 --- a/stage0/stdlib/Init/Notation.c +++ b/stage0/stdlib/Init/Notation.c @@ -72,6 +72,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_guardMsgsOrdering___closed__3; LEAN_EXPORT lean_object* l_term___x3a_x3a__; LEAN_EXPORT lean_object* l_prioMid; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14; static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__19; LEAN_EXPORT lean_object* l_Lean_guardMsgsOrderingArg; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__16; @@ -199,6 +200,7 @@ static lean_object* l_Lean_guardMsgsSpec___closed__5; static lean_object* l_Lean_guardMsgsCmd___closed__12; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_term___x5e_x5e_x5e_____closed__4; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1___closed__2; static lean_object* l_Lean_rawStx_quot___closed__9; static lean_object* l_Lean_byElab___closed__2; @@ -299,6 +301,7 @@ static lean_object* l_termWithout__expected__type_____closed__5; static lean_object* l_term___x5e_____closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__10; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____2(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___u2218____1___closed__9; static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__8; static lean_object* l_stx___x2a___closed__2; @@ -308,6 +311,7 @@ static lean_object* l_term_x7e_x7e_x7e_____closed__1; static lean_object* l_Lean_guardMsgsCmd___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__23; static lean_object* l_term___x5e_x5e_x5e_____closed__6; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3e____1___closed__6; static lean_object* l_Lean_byElab___closed__7; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__16; @@ -323,6 +327,8 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term_x2d____2( static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1___closed__4; static lean_object* l_termIfThenElse___closed__8; static lean_object* l___aux__Init__Notation______macroRules__term___x5c_x2f____1___closed__1; +static lean_object* l_Lean_Parser_commandSeal_______closed__2; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1; static lean_object* l_term___u2264_____closed__6; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__15; static lean_object* l_term___x3d_x3d_____closed__2; @@ -358,6 +364,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1 static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1___closed__5; static lean_object* l_term___x5e_x5e_x5e_____closed__2; static lean_object* l_Lean_builtin__command__code__action___closed__3; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3; static lean_object* l_Lean_command__code__action___closed__6; static lean_object* l_Lean_runMeta___closed__1; static lean_object* l_Lean_guardMsgsCmd___closed__14; @@ -370,6 +377,7 @@ static lean_object* l_Lean_Parser_checkTactic___closed__8; LEAN_EXPORT lean_object* l_term___x5e_x5e_x5e__; static lean_object* l_Lean_guardMsgsOrdering___closed__4; static lean_object* l_term___x3d_____closed__3; +static lean_object* l_Lean_Parser_commandUnseal_______closed__6; static lean_object* l_prec_x28___x29___closed__10; LEAN_EXPORT lean_object* l_Lean_guardMsgsSpecElt; static lean_object* l_Lean_guardMsgsOrderingArg___closed__2; @@ -443,6 +451,7 @@ static lean_object* l_Lean_Parser_Syntax_subPrec___closed__5; static lean_object* l_termDepIfThenElse___closed__17; static lean_object* l_Lean_Parser_Syntax_addPrio___closed__2; static lean_object* l_prioDefault___closed__2; +static lean_object* l_Lean_Parser_commandSeal_______closed__1; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____2(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term_x7e_x7e_x7e____1___closed__1; @@ -456,6 +465,7 @@ static lean_object* l_term___x3c_x3c_x3c_____closed__1; static lean_object* l_termDepIfThenElse___closed__18; static lean_object* l_Lean_guardMsgsCmd___closed__18; static lean_object* l_termMax__prec___closed__1; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12; static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__5; LEAN_EXPORT lean_object* l_term___x3c_x2a__; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__6; @@ -469,7 +479,9 @@ static lean_object* l_term___x2d_____closed__3; static lean_object* l_term___x3c_x7c_____closed__1; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__1; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15; static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26____1___closed__5; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__8; static lean_object* l_Lean_guardMsgsOrderingArg___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__4; @@ -478,6 +490,7 @@ static lean_object* l_Lean_guardMsgsWhitespace___closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x3c____1___closed__3; static lean_object* l_Lean_guardMsgsSpec___closed__4; static lean_object* l_stx___x3f___closed__4; +static lean_object* l_Lean_Parser_commandUnseal_______closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HMul__hMul__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Category_level; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____1(lean_object*, lean_object*, lean_object*); @@ -586,6 +599,7 @@ static lean_object* l_Lean_deprecated___closed__18; 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___aux__Init__Notation______unexpand__Prod__1(lean_object*, lean_object*, lean_object*); static lean_object* l_termIfThenElse___closed__9; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term_x2d____2___closed__1; static lean_object* l_Lean_guardMsgsWhitespaceArg___closed__12; @@ -594,6 +608,7 @@ static lean_object* l___aux__Init__Notation______macroRules__termIfLet__1___clos static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26____1___closed__7; static lean_object* l_term_x7e_x7e_x7e_____closed__3; static lean_object* l___aux__Init__Notation______macroRules__term_x21____1___closed__7; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9; static lean_object* l_term___x24_______closed__10; static lean_object* l_termDepIfThenElse___closed__1; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__stx___x3c_x7c_x3e____1(lean_object*, lean_object*, lean_object*); @@ -620,6 +635,7 @@ LEAN_EXPORT lean_object* l_Lean_guardMsgsCmd; static lean_object* l_boolIfThenElse___closed__13; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x2b____2(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3e_____closed__6; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8; static lean_object* l_stx___x2c_x2a___closed__2; static lean_object* l_term___x2f_x5c_____closed__5; LEAN_EXPORT lean_object* l_stx_x21__; @@ -672,6 +688,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___u2208____1__ LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HShiftRight__hShiftRight__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termWithout__expected__type____1___closed__4; static lean_object* l_term___u2265_____closed__2; +static lean_object* l_Lean_Parser_commandUnseal_______closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___u2223____1___closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__termIfThenElse__1(lean_object*, lean_object*, lean_object*); static lean_object* l_stx___x3c_x7c_x3e_____closed__8; @@ -685,6 +702,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x5e____ static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1___closed__1; static lean_object* l_term___u2209_____closed__3; LEAN_EXPORT lean_object* l_term___x2d__; +static lean_object* l_Lean_Parser_commandSeal_______closed__6; static lean_object* l_Lean_guardMsgsWhitespaceArg___closed__7; static lean_object* l_term___x2f_____closed__1; static lean_object* l___aux__Init__Notation______macroRules__term___x5e_x5e_x5e____1___closed__7; @@ -697,6 +715,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__termMax__prec_ static lean_object* l_stx___x3c_x7c_x3e_____closed__6; static lean_object* l_boolIfThenElse___closed__7; static lean_object* l_term___x2f_____closed__5; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2; static lean_object* l_Lean_guardMsgsFilter___closed__6; static lean_object* l_prioMid___closed__5; LEAN_EXPORT lean_object* l_Lean_Attr_coe; @@ -720,6 +739,7 @@ static lean_object* l_boolIfThenElse___closed__10; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__HSub__hSub__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termIfLet__1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term_x7b___x3a___x2f_x2f___x7d__1___closed__5; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16; static lean_object* l_precMin1___closed__2; static lean_object* l_termIfThenElse___closed__11; static lean_object* l_prec_x28___x29___closed__5; @@ -737,6 +757,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26_x26_ static lean_object* l_Lean_rawStx_quot___closed__10; static lean_object* l_termDepIfThenElse___closed__14; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3e____1___closed__2; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2; static lean_object* l_term_x21_____closed__2; static lean_object* l_term___xd7_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x7c_x3e____1___closed__9; @@ -774,6 +795,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c_x7c_ static lean_object* l_stx_x21_____closed__3; LEAN_EXPORT lean_object* l_stx___x2c_x2b; static lean_object* l_Lean_Parser_Syntax_addPrec___closed__7; +LEAN_EXPORT lean_object* l_Lean_Parser_commandSeal____; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3c_x3c____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__prioMid__1___closed__1; static lean_object* l_Lean_byElab___closed__3; @@ -840,6 +862,7 @@ static lean_object* l_Lean_guardMsgsWhitespaceArg___closed__14; static lean_object* l_stx___x3c_x7c_x3e_____closed__3; static lean_object* l_term_x7b___x3a___x2f_x2f___x7d___closed__16; static lean_object* l___aux__Init__Notation______macroRules__termIfLet__1___closed__7; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10; static lean_object* l_term___x3e_x3e_x3e_____closed__2; static lean_object* l_term___x3e_x3e_____closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3e____1___closed__1; @@ -903,6 +926,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d____2 static lean_object* l_term___x5e_____closed__4; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__13; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__GT__gt__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_commandSeal_______closed__7; static lean_object* l_term___u2209_____closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__stx___x2a__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__stx___x2b__1___closed__2; @@ -1046,6 +1070,7 @@ static lean_object* l_Lean_Parser_Tactic_caseArg___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__SeqLeft__seqLeft__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_builtin__command__code__action; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__19; +static lean_object* l_Lean_Parser_commandUnseal_______closed__4; static lean_object* l_term___xd7_____closed__7; static lean_object* l___aux__Init__Notation______macroRules__term_x2d____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term_xac____1___closed__5; @@ -1056,6 +1081,7 @@ LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3d_x3d static lean_object* l___aux__Init__Notation______macroRules__term___x26_x26____1___closed__4; static lean_object* l_Lean_Parser_checkTactic___closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Dvd__dvd__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13; static lean_object* l_term___x5c_x2f_____closed__6; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Membership__mem__1(lean_object*, lean_object*, lean_object*); static lean_object* l_term___x3c_x2a_x3e_____closed__2; @@ -1089,6 +1115,7 @@ LEAN_EXPORT lean_object* l_term___x5e__; static lean_object* l_term___x3a_x3a_____closed__6; static lean_object* l_Lean_guardMsgsOrderingArg___closed__1; LEAN_EXPORT lean_object* l_Lean_instCoeSyntaxNodeKindSyntaxNodeKinds(lean_object*); +static lean_object* l_Lean_Parser_commandSeal_______closed__5; static lean_object* l_term___x3e_x3e_x3d_____closed__1; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__22; static lean_object* l_precArg___closed__2; @@ -1101,6 +1128,7 @@ static lean_object* l___aux__Init__Notation______macroRules__termIfThenElse__1__ static lean_object* l_term___xd7_____closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___u2208____1___closed__3; static lean_object* l___aux__Init__Notation______macroRules__term___x3e_x3d____1___closed__1; +static lean_object* l_Lean_Parser_commandUnseal_______closed__1; static lean_object* l_Lean_Parser_Syntax_subPrec___closed__6; static lean_object* l_Lean_withAnnotateTerm___closed__4; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1(lean_object*, lean_object*, lean_object*); @@ -1246,10 +1274,12 @@ static lean_object* l_Lean_guardMsgsWhitespace___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x3d____1___closed__2; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__BEq__beq__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_Tactic_caseArg; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7; static lean_object* l___aux__Init__Notation______macroRules__term___x7c_x7c_x7c____1___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___u2223____1___closed__3; LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Functor__map__1(lean_object*, lean_object*, lean_object*); static lean_object* l_termDepIfThenElse___closed__10; +LEAN_EXPORT lean_object* l_Lean_Parser_commandUnseal____; LEAN_EXPORT lean_object* l___aux__Init__Notation______macroRules__term___x3c____1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__LT__lt__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__Bind__bind__1___boxed(lean_object*, lean_object*, lean_object*); @@ -1288,6 +1318,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a____1 static lean_object* l_Lean_rawStx_quot___closed__6; static lean_object* l_boolIfThenElse___closed__5; LEAN_EXPORT lean_object* l_stx___x2c_x2a_x2c_x3f; +static lean_object* l_Lean_Parser_commandUnseal_______closed__3; static lean_object* l_termIfThenElse___closed__10; static lean_object* l_stx___x2c_x2b_x2c_x3f___closed__4; static lean_object* l___aux__Init__Notation______macroRules__term_x2d____1___closed__6; @@ -1330,6 +1361,7 @@ static lean_object* l_Lean_Parser_checkTactic___closed__12; static lean_object* l___aux__Init__Notation______macroRules__termDepIfThenElse__1___closed__3; static lean_object* l_Lean_Parser_checkTacticFailure___closed__2; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__20; +static lean_object* l_Lean_Parser_commandSeal_______closed__3; static lean_object* l_term___u2227_____closed__5; static lean_object* l_term___x7c_x7c_x7c_____closed__7; static lean_object* l_termDepIfThenElse___closed__26; @@ -1411,6 +1443,7 @@ static lean_object* l_termDepIfThenElse___closed__4; static lean_object* l_term___x3e_x3e_x3d_____closed__5; static lean_object* l_term___x3c_x7c_x3e_____closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___u2209____1___closed__3; +static lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6; static lean_object* l_Lean_deprecated___closed__2; static lean_object* l_Lean_guardMsgsCmd___closed__8; static lean_object* l_termIfLet___closed__13; @@ -1438,6 +1471,7 @@ static lean_object* l_termIfLet___closed__3; static lean_object* l_term___x26_x26_x26_____closed__4; static lean_object* l_term___x3d_x3d_____closed__1; static lean_object* l_term___x2a_x3e_____closed__6; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_binderIdent___closed__6; static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a_x3e____1___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term___x2a_x3e____1___closed__3; @@ -1478,6 +1512,7 @@ static lean_object* l___aux__Init__Notation______macroRules__term___x3c_x2a____1 LEAN_EXPORT lean_object* l___aux__Init__Notation______unexpand__SeqRight__seqRight__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withAnnotateTerm___closed__3; static lean_object* l_term___x24_______closed__4; +static lean_object* l_Lean_Parser_commandSeal_______closed__4; static lean_object* l_Lean_guardMsgsFilterSeverity___closed__5; static lean_object* l_Lean_runElab___closed__5; static lean_object* l___aux__Init__Notation______macroRules__term_x7e_x7e_x7e____1___closed__7; @@ -1494,6 +1529,7 @@ static lean_object* l_stx___x2c_x2a_x2c_x3f___closed__2; static lean_object* l___aux__Init__Notation______macroRules__term___x2b_x2b____1___closed__6; static lean_object* l___aux__Init__Notation______macroRules__stx___x2b__1___closed__7; static lean_object* l_Lean_builtin__command__code__action___closed__4; +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__Notation______macroRules__termIfThenElse__1___closed__7; LEAN_EXPORT lean_object* l_term_x21__; static lean_object* l_stx___x2a___closed__1; @@ -25787,6 +25823,550 @@ x_1 = l_Lean_Parser_checkSimpFailure___closed__6; return x_1; } } +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("commandSeal__", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l_Lean_Parser_commandSeal_______closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("seal ", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_commandSeal_______closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__5; +x_2 = l_Lean_deprecated___closed__4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; +x_2 = l_Lean_Parser_commandSeal_______closed__4; +x_3 = l_Lean_Parser_commandSeal_______closed__5; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal_______closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_commandSeal_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_commandSeal_______closed__6; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandSeal____() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_commandSeal_______closed__7; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("commandUnseal__", 15); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l_Lean_Parser_commandUnseal_______closed__1; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unseal ", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser_commandUnseal_______closed__3; +x_2 = lean_alloc_ctor(5, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__7; +x_2 = l_Lean_Parser_commandUnseal_______closed__4; +x_3 = l_Lean_Parser_commandSeal_______closed__5; +x_4 = lean_alloc_ctor(2, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal_______closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Lean_Parser_commandUnseal_______closed__2; +x_2 = lean_unsigned_to_nat(1022u); +x_3 = l_Lean_Parser_commandUnseal_______closed__5; +x_4 = lean_alloc_ctor(3, 3, 0); +lean_ctor_set(x_4, 0, x_1); +lean_ctor_set(x_4, 1, x_2); +lean_ctor_set(x_4, 2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Parser_commandUnseal____() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Parser_commandUnseal_______closed__6; +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___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_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("attribute", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__1; +x_4 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______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_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("[", 1); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("attrInstance", 12); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__1; +x_4 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("attrKind", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______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_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__1; +x_4 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("local", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l___aux__Init__Notation______macroRules__term___u2218____1___closed__1; +x_4 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("simple", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______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_Parser_Syntax_addPrec___closed__1; +x_2 = l_Lean_Parser_Syntax_addPrec___closed__2; +x_3 = l_Lean_Attr_coe___closed__1; +x_4 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("irreducible", 11); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("]", 1); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Parser_commandSeal_______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; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l_Lean_Syntax_getArgs(x_9); +lean_dec(x_9); +x_11 = lean_ctor_get(x_2, 5); +lean_inc(x_11); +x_12 = 0; +x_13 = l_Lean_SourceInfo_fromRef(x_11, x_12); +x_14 = lean_ctor_get(x_2, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2; +lean_inc(x_13); +x_17 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4; +lean_inc(x_13); +x_19 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9; +lean_inc(x_13); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10; +lean_inc(x_13); +x_23 = l_Lean_Syntax_node1(x_13, x_22, x_21); +x_24 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__7; +lean_inc(x_13); +x_25 = l_Lean_Syntax_node1(x_13, x_24, x_23); +x_26 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8; +lean_inc(x_13); +x_27 = l_Lean_Syntax_node1(x_13, x_26, x_25); +x_28 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15; +x_29 = l_Lean_addMacroScope(x_15, x_28, x_14); +x_30 = lean_box(0); +x_31 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14; +lean_inc(x_13); +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_30); +x_33 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +lean_inc(x_13); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_13); +lean_ctor_set(x_34, 1, x_24); +lean_ctor_set(x_34, 2, x_33); +x_35 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12; +lean_inc(x_13); +x_36 = l_Lean_Syntax_node2(x_13, x_35, x_32, x_34); +x_37 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6; +lean_inc(x_13); +x_38 = l_Lean_Syntax_node2(x_13, x_37, x_27, x_36); +lean_inc(x_13); +x_39 = l_Lean_Syntax_node1(x_13, x_24, x_38); +x_40 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16; +lean_inc(x_13); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_13); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Array_appendCore___rarg(x_33, x_10); +lean_dec(x_10); +lean_inc(x_13); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_13); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_42); +x_44 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3; +x_45 = l_Lean_Syntax_node5(x_13, x_44, x_17, x_19, x_39, x_41, x_43); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_3); +return x_46; +} +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("semireducible", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1; +x_2 = l_String_toSubstring_x27(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; uint8_t x_5; +x_4 = l_Lean_Parser_commandUnseal_______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; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; 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_8 = lean_unsigned_to_nat(1u); +x_9 = l_Lean_Syntax_getArg(x_1, x_8); +lean_dec(x_1); +x_10 = l_Lean_Syntax_getArgs(x_9); +lean_dec(x_9); +x_11 = lean_ctor_get(x_2, 5); +lean_inc(x_11); +x_12 = 0; +x_13 = l_Lean_SourceInfo_fromRef(x_11, x_12); +x_14 = lean_ctor_get(x_2, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2; +lean_inc(x_13); +x_17 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_17, 0, x_13); +lean_ctor_set(x_17, 1, x_16); +x_18 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4; +lean_inc(x_13); +x_19 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +x_20 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9; +lean_inc(x_13); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_13); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10; +lean_inc(x_13); +x_23 = l_Lean_Syntax_node1(x_13, x_22, x_21); +x_24 = l___aux__Init__Notation______macroRules__stx___x2b__1___closed__7; +lean_inc(x_13); +x_25 = l_Lean_Syntax_node1(x_13, x_24, x_23); +x_26 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8; +lean_inc(x_13); +x_27 = l_Lean_Syntax_node1(x_13, x_26, x_25); +x_28 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3; +x_29 = l_Lean_addMacroScope(x_15, x_28, x_14); +x_30 = lean_box(0); +x_31 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2; +lean_inc(x_13); +x_32 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_32, 0, x_13); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set(x_32, 2, x_29); +lean_ctor_set(x_32, 3, x_30); +x_33 = l___aux__Init__Notation______macroRules__stx___x2c_x2a__1___closed__10; +lean_inc(x_13); +x_34 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_34, 0, x_13); +lean_ctor_set(x_34, 1, x_24); +lean_ctor_set(x_34, 2, x_33); +x_35 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12; +lean_inc(x_13); +x_36 = l_Lean_Syntax_node2(x_13, x_35, x_32, x_34); +x_37 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6; +lean_inc(x_13); +x_38 = l_Lean_Syntax_node2(x_13, x_37, x_27, x_36); +lean_inc(x_13); +x_39 = l_Lean_Syntax_node1(x_13, x_24, x_38); +x_40 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16; +lean_inc(x_13); +x_41 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_41, 0, x_13); +lean_ctor_set(x_41, 1, x_40); +x_42 = l_Array_appendCore___rarg(x_33, x_10); +lean_dec(x_10); +lean_inc(x_13); +x_43 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_43, 0, x_13); +lean_ctor_set(x_43, 1, x_24); +lean_ctor_set(x_43, 2, x_42); +x_44 = l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3; +x_45 = l_Lean_Syntax_node5(x_13, x_44, x_17, x_19, x_39, x_41, x_43); +x_46 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set(x_46, 1, x_3); +return x_46; +} +} +} lean_object* initialize_Init_Prelude(uint8_t builtin, lean_object*); lean_object* initialize_Init_Coe(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -28504,6 +29084,74 @@ l_Lean_Parser_checkSimpFailure___closed__6 = _init_l_Lean_Parser_checkSimpFailur lean_mark_persistent(l_Lean_Parser_checkSimpFailure___closed__6); l_Lean_Parser_checkSimpFailure = _init_l_Lean_Parser_checkSimpFailure(); lean_mark_persistent(l_Lean_Parser_checkSimpFailure); +l_Lean_Parser_commandSeal_______closed__1 = _init_l_Lean_Parser_commandSeal_______closed__1(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__1); +l_Lean_Parser_commandSeal_______closed__2 = _init_l_Lean_Parser_commandSeal_______closed__2(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__2); +l_Lean_Parser_commandSeal_______closed__3 = _init_l_Lean_Parser_commandSeal_______closed__3(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__3); +l_Lean_Parser_commandSeal_______closed__4 = _init_l_Lean_Parser_commandSeal_______closed__4(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__4); +l_Lean_Parser_commandSeal_______closed__5 = _init_l_Lean_Parser_commandSeal_______closed__5(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__5); +l_Lean_Parser_commandSeal_______closed__6 = _init_l_Lean_Parser_commandSeal_______closed__6(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__6); +l_Lean_Parser_commandSeal_______closed__7 = _init_l_Lean_Parser_commandSeal_______closed__7(); +lean_mark_persistent(l_Lean_Parser_commandSeal_______closed__7); +l_Lean_Parser_commandSeal____ = _init_l_Lean_Parser_commandSeal____(); +lean_mark_persistent(l_Lean_Parser_commandSeal____); +l_Lean_Parser_commandUnseal_______closed__1 = _init_l_Lean_Parser_commandUnseal_______closed__1(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__1); +l_Lean_Parser_commandUnseal_______closed__2 = _init_l_Lean_Parser_commandUnseal_______closed__2(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__2); +l_Lean_Parser_commandUnseal_______closed__3 = _init_l_Lean_Parser_commandUnseal_______closed__3(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__3); +l_Lean_Parser_commandUnseal_______closed__4 = _init_l_Lean_Parser_commandUnseal_______closed__4(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__4); +l_Lean_Parser_commandUnseal_______closed__5 = _init_l_Lean_Parser_commandUnseal_______closed__5(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__5); +l_Lean_Parser_commandUnseal_______closed__6 = _init_l_Lean_Parser_commandUnseal_______closed__6(); +lean_mark_persistent(l_Lean_Parser_commandUnseal_______closed__6); +l_Lean_Parser_commandUnseal____ = _init_l_Lean_Parser_commandUnseal____(); +lean_mark_persistent(l_Lean_Parser_commandUnseal____); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__1 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__1(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__1); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__2); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__3); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__4); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__5); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__6); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__7); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__8); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__9); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__10); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__11); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__12); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__13); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__14); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__15); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandSeal______1___closed__16); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__1); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__2); +l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3 = _init_l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3(); +lean_mark_persistent(l_Lean_Parser___aux__Init__Notation______macroRules__Lean__Parser__commandUnseal______1___closed__3); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/Prelude.c b/stage0/stdlib/Init/Prelude.c index 2aa543d39ddd..54093034feab 100644 --- a/stage0/stdlib/Init/Prelude.c +++ b/stage0/stdlib/Init/Prelude.c @@ -10803,7 +10803,7 @@ static lean_object* _init_l_Lean_maxRecDepthErrorMessage___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("maximum recursion depth has been reached (use `set_option maxRecDepth ` to increase limit)", 95); +x_1 = lean_mk_string_from_bytes("maximum recursion depth has been reached\nuse `set_option maxRecDepth ` to increase limit\nuse `set_option diagnostics true` to get diagnostic information", 157); return x_1; } } diff --git a/stage0/stdlib/Init/System/IO.c b/stage0/stdlib/Init/System/IO.c index 0642659237ab..9e2b0066a607 100644 --- a/stage0/stdlib/Init/System/IO.c +++ b/stage0/stdlib/Init/System/IO.c @@ -311,6 +311,7 @@ static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202__ LEAN_EXPORT lean_object* l_IO_sleep(uint32_t, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_readToEnd_loop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_removeDirAll___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_IO_Process_getCurrentDir___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240_(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_timeit___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -411,6 +412,7 @@ static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init LEAN_EXPORT lean_object* l_IO_FS_instReprDirEntry; LEAN_EXPORT lean_object* l_IO_FS_Mode_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l_IO_FS_instBEqFileType___closed__1; +lean_object* lean_io_process_set_current_dir(lean_object*, lean_object*); static lean_object* l_IO_FS_Mode_noConfusion___rarg___closed__1; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; lean_object* lean_io_prim_handle_get_line(lean_object*, lean_object*); @@ -487,6 +489,7 @@ LEAN_EXPORT lean_object* l_IO_FS_Handle_putStr___boxed(lean_object*, lean_object LEAN_EXPORT lean_object* l_IO_bindTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_bindTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_IO_Process_setCurrentDir___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_mapTasks___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__2; LEAN_EXPORT lean_object* l_Lean_instEvalBaseIO___rarg(lean_object*, lean_object*, uint8_t, lean_object*); @@ -698,6 +701,7 @@ lean_object* lean_io_create_dir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_putStrLn___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4; LEAN_EXPORT lean_object* l_EIO_catchExceptions(lean_object*, lean_object*); +lean_object* lean_io_process_get_current_dir(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instLTSystemTime; static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5; LEAN_EXPORT lean_object* l_Lean_instEvalOfRepr___rarg(lean_object*, lean_object*, uint8_t, lean_object*); @@ -7321,6 +7325,23 @@ lean_dec(x_1); return x_3; } } +LEAN_EXPORT lean_object* l_IO_Process_getCurrentDir___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_io_process_get_current_dir(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_IO_Process_setCurrentDir___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_io_process_set_current_dir(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} LEAN_EXPORT lean_object* l_IO_Process_getPID___boxed(lean_object* x_1) { _start: { diff --git a/stage0/stdlib/Init/Tactics.c b/stage0/stdlib/Init/Tactics.c index e67a7b05c5a5..c6531641dc2a 100644 --- a/stage0/stdlib/Init/Tactics.c +++ b/stage0/stdlib/Init/Tactics.c @@ -6344,34 +6344,77 @@ return x_7; } else { -lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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_8 = lean_ctor_get(x_2, 5); lean_inc(x_8); lean_dec(x_2); x_9 = 0; x_10 = l_Lean_SourceInfo_fromRef(x_8, x_9); -x_11 = l_Lean_Parser_Tactic_fail___closed__1; +x_11 = l_Lean_Parser_Tactic_case_x27___closed__1; lean_inc(x_10); x_12 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRfl__1___closed__1; +x_13 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExfalso__1___closed__19; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_Parser_Tactic_traceMessage___closed__6; +x_15 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExfalso__1___closed__18; lean_inc(x_10); x_16 = l_Lean_Syntax_node1(x_10, x_15, x_14); -x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExfalso__1___closed__14; +x_17 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__2; lean_inc(x_10); x_18 = l_Lean_Syntax_node1(x_10, x_17, x_16); -x_19 = l_Lean_Parser_Tactic_fail___closed__2; -x_20 = l_Lean_Syntax_node2(x_10, x_19, x_12, 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_3); -return x_21; +x_19 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticExfalso__1___closed__14; +x_20 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__5; +lean_inc(x_10); +x_21 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_20); +x_22 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__4; +lean_inc(x_10); +x_23 = l_Lean_Syntax_node2(x_10, x_22, x_18, x_21); +lean_inc(x_10); +x_24 = l_Lean_Syntax_node1(x_10, x_19, x_23); +x_25 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticNext___x3d_x3e____1___closed__6; +lean_inc(x_10); +x_26 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_26, 0, x_10); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Lean_Parser_Tactic_fail___closed__1; +lean_inc(x_10); +x_28 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_28, 0, x_10); +lean_ctor_set(x_28, 1, x_27); +x_29 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticRfl__1___closed__1; +lean_inc(x_10); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_10); +lean_ctor_set(x_30, 1, x_29); +x_31 = l_Lean_Parser_Tactic_traceMessage___closed__6; +lean_inc(x_10); +x_32 = l_Lean_Syntax_node1(x_10, x_31, x_30); +lean_inc(x_10); +x_33 = l_Lean_Syntax_node1(x_10, x_19, x_32); +x_34 = l_Lean_Parser_Tactic_fail___closed__2; +lean_inc(x_10); +x_35 = l_Lean_Syntax_node2(x_10, x_34, x_28, x_33); +lean_inc(x_10); +x_36 = l_Lean_Syntax_node1(x_10, x_19, x_35); +x_37 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__3; +lean_inc(x_10); +x_38 = l_Lean_Syntax_node1(x_10, x_37, x_36); +x_39 = l_Lean_Parser_Tactic___aux__Init__Tactics______macroRules__Lean__Parser__Tactic__tacticTry____1___closed__1; +lean_inc(x_10); +x_40 = l_Lean_Syntax_node1(x_10, x_39, x_38); +x_41 = l_Lean_Parser_Tactic_case_x27___closed__2; +x_42 = l_Lean_Syntax_node4(x_10, x_41, x_12, x_24, x_26, x_40); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_3); +return x_43; } } } diff --git a/stage0/stdlib/Init/WFTactics.c b/stage0/stdlib/Init/WFTactics.c index 329a6c415ded..8ae438a6893d 100644 --- a/stage0/stdlib/Init/WFTactics.c +++ b/stage0/stdlib/Init/WFTactics.c @@ -13,6 +13,8 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_tacticDecreasing__trivial__pre__omega___closed__1; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__61; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__62; @@ -28,9 +30,7 @@ static lean_object* l_tacticDecreasing__with_____closed__10; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__29; static lean_object* l_tacticDecreasing__tactic___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__12; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6; static lean_object* l_tacticSimp__wf___closed__4; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__21; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__83; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__101; @@ -39,6 +39,7 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1_ static lean_object* l_tacticDecreasing__with_____closed__8; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__40; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__65; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__10; @@ -47,24 +48,27 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__ static lean_object* l_tacticDecreasing__with_____closed__7; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__3; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__20; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__74; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__14; static lean_object* l_tacticDecreasing__tactic___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__19; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__24; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__31; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__27; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__48; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__32; static lean_object* l_tacticDecreasing__trivial___closed__4; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3; +LEAN_EXPORT lean_object* l_tacticDecreasing__trivial__pre__omega; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__49; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__3; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__75; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__37; -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1; static lean_object* l_tacticDecreasing__with_____closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__24; @@ -74,15 +78,17 @@ lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_ uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__97; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__52; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4; LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__16; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__18; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__81; static lean_object* l_tacticDecreasing__trivial___closed__2; -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__33; +static lean_object* l_tacticDecreasing__trivial__pre__omega___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__50; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__102; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__88; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__71; @@ -92,9 +98,11 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1_ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__17; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__8; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__7; LEAN_EXPORT lean_object* l_tacticDecreasing__trivial; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__85; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__41; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__98; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__79; @@ -114,15 +122,12 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__ static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__9; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__82; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__38; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__44; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__12; static lean_object* l_tacticSimp__wf___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__30; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__54; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__58; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__6; LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__10; @@ -132,30 +137,29 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__ static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__18; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__23; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__10; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2; +static lean_object* l_tacticDecreasing__trivial__pre__omega___closed__2; LEAN_EXPORT lean_object* l_tacticDecreasing__tactic; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__13; static lean_object* l_tacticDecreasing__trivial___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__94; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__51; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__87; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__76; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__47; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_tacticDecreasing__with_____closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__4; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3; LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3(lean_object*, lean_object*, lean_object*); -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__66; static lean_object* l_tacticDecreasing__tactic___closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__95; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__89; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__53; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__2___closed__2; @@ -170,60 +174,62 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1_ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__99; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__21; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__26; +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__8; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__19; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1___closed__3; +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__28; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__103; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1___closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__22; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__11; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__96; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__45; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__92; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__6; +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__7; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__18; lean_object* l_Lean_Syntax_node7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__91; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__2(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__7; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__42; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__3; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__9; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__31; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__25; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__67; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__6; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__8; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__1; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__35; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__2___closed__1; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__12; -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__46; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__5; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8; +static lean_object* l_tacticDecreasing__trivial__pre__omega___closed__3; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__17; static lean_object* l_tacticDecreasing__with_____closed__6; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__73; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__22; +static lean_object* l_tacticDecreasing__trivial__pre__omega___closed__4; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__13; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__43; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__80; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__30; static lean_object* l_tacticDecreasing__trivial___closed__5; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__15; @@ -236,12 +242,14 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__13; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__7; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__68; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__9; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__27; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__14; static lean_object* l_tacticDecreasing__with_____closed__3; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__with____1___closed__5; LEAN_EXPORT lean_object* l_tacticDecreasing__with__; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__33; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__14; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__11; @@ -253,13 +261,11 @@ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1_ static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__77; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__6; static lean_object* l_tacticDecreasing__with_____closed__2; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6; +static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9; static lean_object* l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__64; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3; static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__1___closed__11; LEAN_EXPORT lean_object* l_tacticSimp__wf; static lean_object* l_tacticDecreasing__tactic___closed__3; -static lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5; static lean_object* _init_l_tacticSimp__wf___closed__1() { _start: { @@ -2181,7 +2187,67 @@ return x_15; } } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1() { +static lean_object* _init_l_tacticDecreasing__trivial__pre__omega___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tacticDecreasing_trivial_pre_omega", 34); +return x_1; +} +} +static lean_object* _init_l_tacticDecreasing__trivial__pre__omega___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_tacticDecreasing__trivial__pre__omega___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_tacticDecreasing__trivial__pre__omega___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decreasing_trivial_pre_omega", 28); +return x_1; +} +} +static lean_object* _init_l_tacticDecreasing__trivial__pre__omega___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l_tacticDecreasing__trivial__pre__omega___closed__3; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l_tacticDecreasing__trivial__pre__omega___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_tacticDecreasing__trivial__pre__omega___closed__2; +x_2 = lean_unsigned_to_nat(1024u); +x_3 = l_tacticDecreasing__trivial__pre__omega___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_tacticDecreasing__trivial__pre__omega() { +_start: +{ +lean_object* x_1; +x_1 = l_tacticDecreasing__trivial__pre__omega___closed__5; +return x_1; +} +} +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1() { _start: { lean_object* x_1; @@ -2189,19 +2255,19 @@ x_1 = lean_mk_string_from_bytes("seq1", 4); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__1; x_2 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__2; x_3 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__3; -x_4 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1; +x_4 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3() { _start: { lean_object* x_1; @@ -2209,19 +2275,19 @@ x_1 = lean_mk_string_from_bytes("apply", 5); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__1; x_2 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__2; x_3 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__3; -x_4 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +x_4 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5() { _start: { lean_object* x_1; @@ -2229,16 +2295,16 @@ x_1 = lean_mk_string_from_bytes("Nat.sub_succ_lt_self", 20); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5; +x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7() { _start: { lean_object* x_1; @@ -2246,41 +2312,41 @@ x_1 = lean_mk_string_from_bytes("sub_succ_lt_self", 16); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__88; -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11() { _start: { lean_object* x_1; @@ -2288,11 +2354,11 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__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_tacticDecreasing__trivial___closed__2; +x_4 = l_tacticDecreasing__trivial__pre__omega___closed__2; x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) { @@ -2316,25 +2382,25 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); lean_inc(x_12); lean_dec(x_2); -x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8; +x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8; x_16 = l_Lean_addMacroScope(x_12, x_15, x_11); -x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6; -x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10; +x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6; +x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10; lean_inc(x_10); x_19 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4; +x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4; lean_inc(x_10); x_21 = l_Lean_Syntax_node2(x_10, x_20, x_14, x_19); -x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11; +x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11; lean_inc(x_10); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_10); @@ -2350,7 +2416,7 @@ x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); x_28 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__12; lean_inc(x_10); x_29 = l_Lean_Syntax_node3(x_10, x_28, x_21, x_23, x_27); -x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2; +x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2; x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_31); @@ -2359,7 +2425,7 @@ return x_32; } } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1() { _start: { lean_object* x_1; @@ -2367,16 +2433,16 @@ x_1 = lean_mk_string_from_bytes("Nat.pred_lt'", 12); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1; +x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3() { _start: { lean_object* x_1; @@ -2384,45 +2450,45 @@ x_1 = lean_mk_string_from_bytes("pred_lt'", 8); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__88; -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_tacticDecreasing__trivial___closed__2; +x_4 = l_tacticDecreasing__trivial__pre__omega___closed__2; x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) { @@ -2446,25 +2512,25 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); lean_inc(x_12); lean_dec(x_2); -x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4; +x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4; x_16 = l_Lean_addMacroScope(x_12, x_15, x_11); -x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2; -x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6; +x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2; +x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6; lean_inc(x_10); x_19 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4; +x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4; lean_inc(x_10); x_21 = l_Lean_Syntax_node2(x_10, x_20, x_14, x_19); -x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11; +x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11; lean_inc(x_10); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_10); @@ -2480,7 +2546,7 @@ x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); x_28 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__12; lean_inc(x_10); x_29 = l_Lean_Syntax_node3(x_10, x_28, x_21, x_23, x_27); -x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2; +x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2; x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_31); @@ -2489,7 +2555,7 @@ return x_32; } } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1() { _start: { lean_object* x_1; @@ -2497,16 +2563,16 @@ x_1 = lean_mk_string_from_bytes("Nat.pred_lt", 11); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1; +x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3() { _start: { lean_object* x_1; @@ -2514,67 +2580,67 @@ x_1 = lean_mk_string_from_bytes("pred_lt", 7); return x_1; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__88; -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4; +x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___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__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___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__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8() { +static lean_object* _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5; -x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7; +x_1 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5; +x_2 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__7; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_tacticDecreasing__trivial___closed__2; +x_4 = l_tacticDecreasing__trivial__pre__omega___closed__2; x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) { @@ -2598,25 +2664,25 @@ lean_inc(x_11); x_12 = lean_ctor_get(x_2, 1); lean_inc(x_12); lean_dec(x_2); -x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +x_13 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; lean_inc(x_10); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_10); lean_ctor_set(x_14, 1, x_13); -x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4; +x_15 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4; x_16 = l_Lean_addMacroScope(x_12, x_15, x_11); -x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2; -x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8; +x_17 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2; +x_18 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8; lean_inc(x_10); x_19 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_19, 0, x_10); lean_ctor_set(x_19, 1, x_17); lean_ctor_set(x_19, 2, x_16); lean_ctor_set(x_19, 3, x_18); -x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4; +x_20 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4; lean_inc(x_10); x_21 = l_Lean_Syntax_node2(x_10, x_20, x_14, x_19); -x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11; +x_22 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11; lean_inc(x_10); x_23 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_23, 0, x_10); @@ -2632,7 +2698,7 @@ x_27 = l_Lean_Syntax_node1(x_10, x_26, x_25); x_28 = l___aux__Init__WFTactics______macroRules__tacticSimp__wf__1___closed__12; lean_inc(x_10); x_29 = l_Lean_Syntax_node3(x_10, x_28, x_21, x_23, x_27); -x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2; +x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2; x_31 = l_Lean_Syntax_node1(x_10, x_30, x_29); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_31); @@ -3181,7 +3247,7 @@ lean_inc(x_12); x_29 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_29, 0, x_12); lean_ctor_set(x_29, 1, x_28); -x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3; +x_30 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3; lean_inc(x_12); x_31 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_31, 0, x_12); @@ -3198,7 +3264,7 @@ lean_ctor_set(x_36, 0, x_12); lean_ctor_set(x_36, 1, x_34); lean_ctor_set(x_36, 2, x_33); lean_ctor_set(x_36, 3, x_35); -x_37 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4; +x_37 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4; lean_inc(x_31); lean_inc(x_12); x_38 = l_Lean_Syntax_node2(x_12, x_37, x_31, x_36); @@ -3568,7 +3634,7 @@ lean_ctor_set(x_30, 1, x_29); x_31 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__tactic__1___closed__3; lean_inc(x_10); x_32 = l_Lean_Syntax_node1(x_10, x_31, x_30); -x_33 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11; +x_33 = l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11; lean_inc(x_10); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_10); @@ -3894,56 +3960,68 @@ l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed_ lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__1); l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__2(); lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__3___closed__2); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__1); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__2); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__3); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__4); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__5); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__6); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__7); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__8); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__9); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__10); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__4___closed__11); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__1); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__2); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__3); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__4); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__5); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__5___closed__6); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__1); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__2); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__3); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__4); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__5); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__6); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__7); -l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8(); -lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__6___closed__8); +l_tacticDecreasing__trivial__pre__omega___closed__1 = _init_l_tacticDecreasing__trivial__pre__omega___closed__1(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega___closed__1); +l_tacticDecreasing__trivial__pre__omega___closed__2 = _init_l_tacticDecreasing__trivial__pre__omega___closed__2(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega___closed__2); +l_tacticDecreasing__trivial__pre__omega___closed__3 = _init_l_tacticDecreasing__trivial__pre__omega___closed__3(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega___closed__3); +l_tacticDecreasing__trivial__pre__omega___closed__4 = _init_l_tacticDecreasing__trivial__pre__omega___closed__4(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega___closed__4); +l_tacticDecreasing__trivial__pre__omega___closed__5 = _init_l_tacticDecreasing__trivial__pre__omega___closed__5(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega___closed__5); +l_tacticDecreasing__trivial__pre__omega = _init_l_tacticDecreasing__trivial__pre__omega(); +lean_mark_persistent(l_tacticDecreasing__trivial__pre__omega); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__1); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__2); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__3); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__4); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__5); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__6); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__7); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__8); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__9); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__10); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__1___closed__11); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__1); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__2); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__3); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__4); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__5); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__2___closed__6); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__1); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__2); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__3); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__4); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__5); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__6 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__6(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__6); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__7 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__7(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__7); +l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8 = _init_l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8(); +lean_mark_persistent(l___aux__Init__WFTactics______macroRules__tacticDecreasing__trivial__pre__omega__3___closed__8); l_tacticDecreasing__with_____closed__1 = _init_l_tacticDecreasing__with_____closed__1(); lean_mark_persistent(l_tacticDecreasing__with_____closed__1); l_tacticDecreasing__with_____closed__2 = _init_l_tacticDecreasing__with_____closed__2(); diff --git a/stage0/stdlib/Lean/BuiltinDocAttr.c b/stage0/stdlib/Lean/BuiltinDocAttr.c index 442b7950fa5d..230084d4a71d 100644 --- a/stage0/stdlib/Lean/BuiltinDocAttr.c +++ b/stage0/stdlib/Lean/BuiltinDocAttr.c @@ -663,6 +663,8 @@ x_80 = lean_array_push(x_79, x_78); x_81 = l_Lean_declareBuiltinDocStringAndRanges___lambda__1___closed__6; x_82 = l_Lean_mkAppN(x_81, x_80); x_83 = l_Lean_declareBuiltin(x_17, x_82, x_3, x_4, x_14); +lean_dec(x_4); +lean_dec(x_3); return x_83; } } @@ -761,8 +763,6 @@ x_22 = lean_array_push(x_21, x_19); x_23 = lean_array_push(x_22, x_20); x_24 = l_Lean_declareBuiltinDocStringAndRanges___closed__5; x_25 = l_Lean_mkAppN(x_24, x_23); -lean_inc(x_3); -lean_inc(x_2); x_26 = l_Lean_declareBuiltin(x_18, x_25, x_2, x_3, x_15); if (lean_obj_tag(x_26) == 0) { diff --git a/stage0/stdlib/Lean/Compiler/ExternAttr.c b/stage0/stdlib/Lean/Compiler/ExternAttr.c index e7454d5a2a4a..4e08bdd1cb5f 100644 --- a/stage0/stdlib/Lean/Compiler/ExternAttr.c +++ b/stage0/stdlib/Lean/Compiler/ExternAttr.c @@ -57,11 +57,15 @@ LEAN_EXPORT lean_object* l_Lean_expandExternPatternAux(lean_object*, lean_object lean_object* l_Lean_stringToMessageData(lean_object*); uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____closed__5; +extern lean_object* l_Lean_maxRecDepth; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__28; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_push(lean_object*, uint32_t); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_getExternConstArityExport___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); extern lean_object* l_Lean_projectionFnInfoExt; LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); @@ -72,12 +76,15 @@ size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_getExternConstArityExport___closed__2; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__30; uint8_t l_instDecidableNot___rarg(uint8_t); +static uint8_t l_Lean_getExternConstArityExport___closed__12; +lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getExternEntryFor___boxed(lean_object*, lean_object*); static lean_object* l_Lean_mkSimpleFnCall___closed__3; uint8_t lean_uint32_dec_le(uint32_t, uint32_t); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__27; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__17; static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____closed__4; uint8_t l_String_Iterator_hasNext(lean_object*); @@ -97,6 +104,7 @@ static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstA LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____lambda__3___closed__1; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__24; +static lean_object* l_Lean_getExternConstArityExport___closed__10; LEAN_EXPORT lean_object* l_Lean_mkSimpleFnCall(lean_object*, lean_object*); lean_object* l_List_foldl___at_String_join___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__19; @@ -112,10 +120,13 @@ static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____c LEAN_EXPORT uint8_t l_Lean_isExtern(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____lambda__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_expandExternPattern(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_addExtern___boxed(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* lean_add_extern(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____closed__11; +static lean_object* l_Lean_getExternConstArityExport___lambda__1___closed__1; +extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____lambda__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_getExternEntryForAux___boxed(lean_object*, lean_object*); @@ -160,12 +171,14 @@ LEAN_EXPORT lean_object* l_Lean_ExternEntry_backend(lean_object*); static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___closed__4; extern lean_object* l_Lean_firstFrontendMacroScope; lean_object* l_Lean_ofExcept___at_Lean_initFn____x40_Lean_Class___hyg_768____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getExternConstArityExport___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_expandExternPattern___boxed(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_ExternAttr___hyg_690____closed__7; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_expandExternPatternAux___closed__2; static lean_object* l_Lean_instInhabitedExternAttrData___closed__1; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ExternAttrData_arity_x3f___default; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__21; lean_object* lean_string_append(lean_object*, lean_object*); @@ -182,6 +195,7 @@ static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternA static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__15; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__16; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__18; +static lean_object* l_Lean_getExternConstArityExport___closed__11; static lean_object* l_Lean_getExternConstArityExport___closed__4; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__13; LEAN_EXPORT lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_parseOptNum(lean_object*, lean_object*, lean_object*); @@ -191,6 +205,7 @@ static lean_object* l_Lean_mkSimpleFnCall___closed__2; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__23; static lean_object* l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__8; static lean_object* l_Lean_getExternConstArityExport___closed__9; +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___spec__3___closed__3; static uint8_t l_Lean_expandExternPatternAux___closed__1; static uint8_t l_Lean_expandExternPatternAux___closed__3; @@ -2913,9 +2928,89 @@ lean_dec(x_1); return x_8; } } +static lean_object* _init_l_Lean_getExternConstArityExport___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_getExternConstArityExport___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) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +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_ctor_get(x_5, 4); +lean_dec(x_9); +x_10 = lean_ctor_get(x_5, 2); +lean_dec(x_10); +x_11 = l_Lean_getExternConstArityExport___lambda__1___closed__1; +x_12 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_11); +lean_ctor_set(x_5, 4, x_12); +lean_ctor_set(x_5, 2, x_1); +lean_ctor_set_uint8(x_5, sizeof(void*)*11 + 1, x_2); +x_13 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity(x_3, x_5, x_6, x_7); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_ctor_get(x_5, 3); +x_17 = lean_ctor_get(x_5, 5); +x_18 = lean_ctor_get(x_5, 6); +x_19 = lean_ctor_get(x_5, 7); +x_20 = lean_ctor_get(x_5, 8); +x_21 = lean_ctor_get(x_5, 9); +x_22 = lean_ctor_get(x_5, 10); +x_23 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); +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_24 = l_Lean_getExternConstArityExport___lambda__1___closed__1; +x_25 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_24); +x_26 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_15); +lean_ctor_set(x_26, 2, x_1); +lean_ctor_set(x_26, 3, x_16); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set(x_26, 5, x_17); +lean_ctor_set(x_26, 6, x_18); +lean_ctor_set(x_26, 7, x_19); +lean_ctor_set(x_26, 8, x_20); +lean_ctor_set(x_26, 9, x_21); +lean_ctor_set(x_26, 10, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*11, x_23); +lean_ctor_set_uint8(x_26, sizeof(void*)*11 + 1, x_2); +x_27 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity(x_3, x_26, x_6, x_7); +return x_27; +} +} +} static lean_object* _init_l_Lean_getExternConstArityExport___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("internal exception #", 20); +return x_1; +} +} +static lean_object* _init_l_Lean_getExternConstArityExport___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_expandExternPatternAux___closed__2; x_2 = l___private_Lean_Compiler_ExternAttr_0__Lean_syntaxToExternAttrData___lambda__1___closed__1; @@ -2925,7 +3020,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__2() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -2934,7 +3029,7 @@ x_2 = l_Lean_Core_getMaxHeartbeats(x_1); return x_2; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__3() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -2944,7 +3039,7 @@ x_3 = lean_nat_add(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__4() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__5() { _start: { lean_object* x_1; @@ -2952,21 +3047,21 @@ x_1 = lean_mk_string_from_bytes("_uniq", 5); return x_1; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__5() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_getExternConstArityExport___closed__4; +x_2 = l_Lean_getExternConstArityExport___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__6() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_getExternConstArityExport___closed__5; +x_1 = l_Lean_getExternConstArityExport___closed__6; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2974,7 +3069,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__7() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__8() { _start: { lean_object* x_1; lean_object* x_2; @@ -2985,7 +3080,7 @@ lean_ctor_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__8() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__9() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; @@ -2999,7 +3094,7 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); return x_4; } } -static lean_object* _init_l_Lean_getExternConstArityExport___closed__9() { +static lean_object* _init_l_Lean_getExternConstArityExport___closed__10() { _start: { lean_object* x_1; @@ -3007,195 +3102,511 @@ x_1 = lean_mk_string_from_bytes("", 10); return x_1; } } +static lean_object* _init_l_Lean_getExternConstArityExport___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static uint8_t _init_l_Lean_getExternConstArityExport___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; +x_1 = lean_box(0); +x_2 = l_Lean_getExternConstArityExport___closed__11; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); +return x_3; +} +} LEAN_EXPORT lean_object* lean_get_extern_const_arity(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; 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; -x_4 = lean_box(0); -x_5 = l_Lean_getExternConstArityExport___closed__3; -x_6 = l_Lean_getExternConstArityExport___closed__6; -x_7 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__7; -x_8 = l_Lean_getExternConstArityExport___closed__7; -x_9 = l_Lean_getExternConstArityExport___closed__8; -x_10 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_5); -lean_ctor_set(x_10, 2, x_6); -lean_ctor_set(x_10, 3, x_7); -lean_ctor_set(x_10, 4, x_8); -lean_ctor_set(x_10, 5, x_7); -lean_ctor_set(x_10, 6, x_9); -x_11 = lean_io_get_num_heartbeats(x_3); -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_getExternConstArityExport___closed__9; -x_15 = l_Lean_getExternConstArityExport___closed__1; -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_unsigned_to_nat(1000u); -x_18 = lean_box(0); -x_19 = lean_box(0); -x_20 = l_Lean_getExternConstArityExport___closed__2; -x_21 = l_Lean_firstFrontendMacroScope; -x_22 = 0; -x_23 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_23, 0, x_14); -lean_ctor_set(x_23, 1, x_15); -lean_ctor_set(x_23, 2, x_4); -lean_ctor_set(x_23, 3, x_16); -lean_ctor_set(x_23, 4, x_17); -lean_ctor_set(x_23, 5, x_18); -lean_ctor_set(x_23, 6, x_19); -lean_ctor_set(x_23, 7, x_4); -lean_ctor_set(x_23, 8, x_12); -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); -lean_ctor_set_uint8(x_23, sizeof(void*)*11 + 1, x_22); -x_24 = lean_st_mk_ref(x_10, x_13); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_25); -x_27 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity(x_2, x_23, x_25, x_26); -if (lean_obj_tag(x_27) == 0) +lean_object* x_4; lean_object* x_21; lean_object* x_22; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; uint8_t x_71; +x_43 = lean_box(0); +x_44 = l_Lean_getExternConstArityExport___closed__4; +x_45 = l_Lean_getExternConstArityExport___closed__7; +x_46 = l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__7; +x_47 = l_Lean_getExternConstArityExport___closed__8; +x_48 = l_Lean_getExternConstArityExport___closed__9; +x_49 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_49, 0, x_1); +lean_ctor_set(x_49, 1, x_44); +lean_ctor_set(x_49, 2, x_45); +lean_ctor_set(x_49, 3, x_46); +lean_ctor_set(x_49, 4, x_47); +lean_ctor_set(x_49, 5, x_46); +lean_ctor_set(x_49, 6, x_48); +x_50 = lean_io_get_num_heartbeats(x_3); +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_getExternConstArityExport___closed__10; +x_54 = l_Lean_getExternConstArityExport___closed__2; +x_55 = lean_unsigned_to_nat(0u); +x_56 = lean_unsigned_to_nat(1000u); +x_57 = lean_box(0); +x_58 = lean_box(0); +x_59 = l_Lean_getExternConstArityExport___closed__3; +x_60 = l_Lean_firstFrontendMacroScope; +x_61 = 0; +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_53); +lean_ctor_set(x_62, 1, x_54); +lean_ctor_set(x_62, 2, x_43); +lean_ctor_set(x_62, 3, x_55); +lean_ctor_set(x_62, 4, x_56); +lean_ctor_set(x_62, 5, x_57); +lean_ctor_set(x_62, 6, x_58); +lean_ctor_set(x_62, 7, x_43); +lean_ctor_set(x_62, 8, x_51); +lean_ctor_set(x_62, 9, x_59); +lean_ctor_set(x_62, 10, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_61); +x_63 = lean_st_mk_ref(x_49, x_52); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_st_ref_get(x_64, x_65); +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, 0); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_Kernel_isDiagnosticsEnabled(x_69); +lean_dec(x_69); +if (x_70 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_st_ref_get(x_25, x_29); -lean_dec(x_25); -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +uint8_t x_135; +x_135 = l_Lean_getExternConstArityExport___closed__12; +if (x_135 == 0) { -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_30, 0); -lean_dec(x_32); -x_33 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_33, 0, x_28); -lean_ctor_set(x_30, 0, x_33); -return x_30; +uint8_t x_136; +x_136 = 1; +x_71 = x_136; +goto block_134; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_30, 1); -lean_inc(x_34); -lean_dec(x_30); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_28); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_34); -return x_36; +x_71 = x_61; +goto block_134; } } else { -lean_object* x_37; -lean_dec(x_25); -x_37 = lean_ctor_get(x_27, 0); -lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) +uint8_t x_137; +x_137 = l_Lean_getExternConstArityExport___closed__12; +if (x_137 == 0) { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_27, 1); -lean_inc(x_38); -lean_dec(x_27); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = l_Lean_MessageData_toString(x_39, x_38); -if (lean_obj_tag(x_40) == 0) +x_71 = x_61; +goto block_134; +} +else { -uint8_t x_41; -x_41 = !lean_is_exclusive(x_40); -if (x_41 == 0) +uint8_t x_138; +x_138 = 1; +x_71 = x_138; +goto block_134; +} +} +block_20: { -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; +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_4, 0, x_8); +return x_4; } 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; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_4); +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +lean_dec(x_9); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_10); +return x_13; } } else { -uint8_t x_47; -x_47 = !lean_is_exclusive(x_40); -if (x_47 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_4); +if (x_14 == 0) { -lean_object* x_48; lean_object* x_49; -x_48 = lean_ctor_get(x_40, 0); -lean_dec(x_48); -x_49 = lean_box(0); -lean_ctor_set_tag(x_40, 0); -lean_ctor_set(x_40, 0, x_49); -return x_40; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_4, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set_tag(x_4, 0); +lean_ctor_set(x_4, 0, x_16); +return x_4; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_40, 1); -lean_inc(x_50); -lean_dec(x_40); -x_51 = lean_box(0); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -return x_52; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_4, 1); +lean_inc(x_17); +lean_dec(x_4); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; +} +} +} +block_42: +{ +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_MessageData_toString(x_23, x_22); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_tag(x_24, 1); +lean_ctor_set(x_24, 0, x_27); +x_4 = x_24; +goto block_20; +} +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 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_30, 0, x_28); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_4 = x_31; +goto block_20; +} +} +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) +{ +x_4 = x_24; +goto block_20; +} +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); +x_4 = x_35; +goto block_20; } } } else { -uint8_t x_53; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_21, 0); +lean_inc(x_36); +lean_dec(x_21); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_36); +x_38 = l_Lean_getExternConstArityExport___closed__1; +x_39 = lean_string_append(x_38, x_37); lean_dec(x_37); -x_53 = !lean_is_exclusive(x_27); -if (x_53 == 0) +x_40 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_22); +x_4 = x_41; +goto block_20; +} +} +block_134: +{ +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_st_ref_take(x_64, x_68); +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_is_exclusive(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_76 = lean_ctor_get(x_73, 0); +x_77 = lean_ctor_get(x_73, 4); +lean_dec(x_77); +x_78 = l_Lean_getExternConstArityExport___closed__12; +x_79 = l_Lean_Kernel_enableDiag(x_76, x_78); +lean_ctor_set(x_73, 4, x_47); +lean_ctor_set(x_73, 0, x_79); +x_80 = lean_st_ref_set(x_64, x_73, x_74); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_box(0); +lean_inc(x_64); +x_83 = l_Lean_getExternConstArityExport___lambda__1(x_43, x_78, x_2, x_82, x_62, x_64, x_81); +if (lean_obj_tag(x_83) == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_27, 0); -lean_dec(x_54); -x_55 = lean_box(0); -lean_ctor_set_tag(x_27, 0); -lean_ctor_set(x_27, 0, x_55); -return x_27; +lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_64, x_85); +lean_dec(x_64); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_86, 0); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_84); +lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_86, 0, x_89); +x_4 = x_86; +goto block_20; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_86, 0); +x_91 = lean_ctor_get(x_86, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_86); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_90); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +x_4 = x_93; +goto block_20; +} } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_27, 1); -lean_inc(x_56); -lean_dec(x_27); -x_57 = lean_box(0); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_57); -lean_ctor_set(x_58, 1, x_56); -return x_58; +lean_object* x_94; lean_object* x_95; +lean_dec(x_64); +x_94 = lean_ctor_get(x_83, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_83, 1); +lean_inc(x_95); +lean_dec(x_83); +x_21 = x_94; +x_22 = x_95; +goto block_42; +} +} +else +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_96 = lean_ctor_get(x_73, 0); +x_97 = lean_ctor_get(x_73, 1); +x_98 = lean_ctor_get(x_73, 2); +x_99 = lean_ctor_get(x_73, 3); +x_100 = lean_ctor_get(x_73, 5); +x_101 = lean_ctor_get(x_73, 6); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_73); +x_102 = l_Lean_getExternConstArityExport___closed__12; +x_103 = l_Lean_Kernel_enableDiag(x_96, x_102); +x_104 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_97); +lean_ctor_set(x_104, 2, x_98); +lean_ctor_set(x_104, 3, x_99); +lean_ctor_set(x_104, 4, x_47); +lean_ctor_set(x_104, 5, x_100); +lean_ctor_set(x_104, 6, x_101); +x_105 = lean_st_ref_set(x_64, x_104, x_74); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_box(0); +lean_inc(x_64); +x_108 = l_Lean_getExternConstArityExport___lambda__1(x_43, x_102, x_2, x_107, x_62, x_64, x_106); +if (lean_obj_tag(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_object* x_116; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_st_ref_get(x_64, x_110); +lean_dec(x_64); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; +} else { + lean_dec_ref(x_111); + x_114 = lean_box(0); } +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); +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); +x_4 = x_116; +goto block_20; } +else +{ +lean_object* x_117; lean_object* x_118; +lean_dec(x_64); +x_117 = lean_ctor_get(x_108, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_108, 1); +lean_inc(x_118); +lean_dec(x_108); +x_21 = x_117; +x_22 = x_118; +goto block_42; +} +} +} +else +{ +uint8_t x_119; lean_object* x_120; lean_object* x_121; +x_119 = l_Lean_getExternConstArityExport___closed__12; +x_120 = lean_box(0); +lean_inc(x_64); +x_121 = l_Lean_getExternConstArityExport___lambda__1(x_43, x_119, x_2, x_120, x_62, x_64, x_68); +if (lean_obj_tag(x_121) == 0) +{ +lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +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_st_ref_get(x_64, x_123); +lean_dec(x_64); +x_125 = !lean_is_exclusive(x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_124, 0); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_122); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_124, 0, x_127); +x_4 = x_124; +goto block_20; +} +else +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_124, 0); +x_129 = lean_ctor_get(x_124, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_124); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_122); +lean_ctor_set(x_130, 1, x_128); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_129); +x_4 = x_131; +goto block_20; +} +} +else +{ +lean_object* x_132; lean_object* x_133; +lean_dec(x_64); +x_132 = lean_ctor_get(x_121, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_121, 1); +lean_inc(x_133); +lean_dec(x_121); +x_21 = x_132; +x_22 = x_133; +goto block_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_getExternConstArityExport___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_getExternConstArityExport___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +return x_9; } } lean_object* initialize_Init_Data_List_BasicAux(uint8_t builtin, lean_object*); @@ -3354,6 +3765,8 @@ l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__30 = lean_mark_persistent(l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__30); l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__31 = _init_l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__31(); lean_mark_persistent(l___private_Lean_Compiler_ExternAttr_0__Lean_getExternConstArity___closed__31); +l_Lean_getExternConstArityExport___lambda__1___closed__1 = _init_l_Lean_getExternConstArityExport___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_getExternConstArityExport___lambda__1___closed__1); l_Lean_getExternConstArityExport___closed__1 = _init_l_Lean_getExternConstArityExport___closed__1(); lean_mark_persistent(l_Lean_getExternConstArityExport___closed__1); l_Lean_getExternConstArityExport___closed__2 = _init_l_Lean_getExternConstArityExport___closed__2(); @@ -3372,6 +3785,11 @@ l_Lean_getExternConstArityExport___closed__8 = _init_l_Lean_getExternConstArityE lean_mark_persistent(l_Lean_getExternConstArityExport___closed__8); l_Lean_getExternConstArityExport___closed__9 = _init_l_Lean_getExternConstArityExport___closed__9(); lean_mark_persistent(l_Lean_getExternConstArityExport___closed__9); +l_Lean_getExternConstArityExport___closed__10 = _init_l_Lean_getExternConstArityExport___closed__10(); +lean_mark_persistent(l_Lean_getExternConstArityExport___closed__10); +l_Lean_getExternConstArityExport___closed__11 = _init_l_Lean_getExternConstArityExport___closed__11(); +lean_mark_persistent(l_Lean_getExternConstArityExport___closed__11); +l_Lean_getExternConstArityExport___closed__12 = _init_l_Lean_getExternConstArityExport___closed__12(); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/IR/Basic.c b/stage0/stdlib/Lean/Compiler/IR/Basic.c index e847b1797266..5b36526b8024 100644 --- a/stage0/stdlib/Lean/Compiler/IR/Basic.c +++ b/stage0/stdlib/Lean/Compiler/IR/Basic.c @@ -2912,7 +2912,7 @@ x_13 = l_Lean_IR_reshapeAux___closed__2; x_14 = lean_string_append(x_12, x_13); x_15 = l_Lean_IR_reshapeAux___closed__3; x_16 = l_Lean_IR_reshapeAux___closed__4; -x_17 = lean_unsigned_to_nat(128u); +x_17 = lean_unsigned_to_nat(129u); x_18 = lean_unsigned_to_nat(4u); x_19 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_15, x_16, x_17, x_18, x_14); lean_dec(x_14); diff --git a/stage0/stdlib/Lean/Compiler/IR/RC.c b/stage0/stdlib/Lean/Compiler/IR/RC.c index c9832d74cbce..7c25684b9573 100644 --- a/stage0/stdlib/Lean/Compiler/IR/RC.c +++ b/stage0/stdlib/Lean/Compiler/IR/RC.c @@ -257,7 +257,7 @@ static lean_object* _init_l_Lean_IR_ExplicitRC_getDecl___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_IR_ExplicitRC_getDecl___closed__1; x_2 = l_Lean_IR_ExplicitRC_getDecl___closed__2; -x_3 = lean_unsigned_to_nat(35u); +x_3 = lean_unsigned_to_nat(36u); x_4 = lean_unsigned_to_nat(17u); x_5 = l_Lean_IR_ExplicitRC_getDecl___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -358,7 +358,7 @@ static lean_object* _init_l_Lean_IR_ExplicitRC_getVarInfo___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_IR_ExplicitRC_getDecl___closed__1; x_2 = l_Lean_IR_ExplicitRC_getVarInfo___closed__1; -x_3 = lean_unsigned_to_nat(40u); +x_3 = lean_unsigned_to_nat(41u); x_4 = lean_unsigned_to_nat(17u); x_5 = l_Lean_IR_ExplicitRC_getDecl___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -442,7 +442,7 @@ static lean_object* _init_l_Lean_IR_ExplicitRC_getJPParams___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_IR_ExplicitRC_getDecl___closed__1; x_2 = l_Lean_IR_ExplicitRC_getJPParams___closed__1; -x_3 = lean_unsigned_to_nat(45u); +x_3 = lean_unsigned_to_nat(46u); x_4 = lean_unsigned_to_nat(15u); x_5 = l_Lean_IR_ExplicitRC_getDecl___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c index c9fd43c8c640..cdf9916f30a9 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c @@ -14,16 +14,33 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___boxed(lean_object*, lean_object*); +size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___rarg(lean_object*); +uint8_t lean_usize_dec_le(size_t, size_t); uint8_t l_Lean_IR_HasIndex_visitFnBody(lean_object*, lean_object*); +uint64_t lean_uint64_of_nat(lean_object*); +size_t lean_uint64_to_usize(uint64_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2(lean_object*, size_t, lean_object*); +size_t lean_usize_mul(size_t, size_t); uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t l_Lean_IR_CtorInfo_isScalar(lean_object*); +lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); lean_object* l_Lean_IR_FnBody_hasLiveVar(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1___boxed(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_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(lean_object*, uint8_t, size_t, size_t, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_Context_lctx___default; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -34,36 +51,62 @@ LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReus LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse___boxed(lean_object*, lean_object*); lean_object* l_Lean_IR_FnBody_body(lean_object*); +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dfinalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_IR_ResetReuse_R___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1; +lean_object* l_Lean_IR_instBEqVarId___boxed(lean_object*, lean_object*); +static size_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dfinalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___spec__1(lean_object*, lean_object*, size_t, size_t); +lean_object* lean_usize_to_nat(size_t); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object*); uint8_t l_Lean_IR_FnBody_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_IR_instHashableVarId___boxed(lean_object*); +static lean_object* l_Lean_IR_Decl_insertResetReuse___closed__1; lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3; LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_IR_MaxIndex_collectDecl(lean_object*, lean_object*); +size_t lean_usize_sub(size_t, size_t); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); +size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_IR_LocalContext_addJP(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(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_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_Context_casesVars___default; LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); +size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object* x_1, lean_object* x_2) { _start: { @@ -539,6 +582,76 @@ lean_dec(x_2); return x_4; } } +static lean_object* _init_l_Lean_IR_ResetReuse_Context_lctx___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1() { +_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_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_instBEqVarId___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_IR_instHashableVarId___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2; +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_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5; +return x_1; +} +} +static lean_object* _init_l_Lean_IR_ResetReuse_Context_casesVars___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; +return x_1; +} +} LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___rarg(lean_object* x_1) { _start: { @@ -1028,628 +1141,644 @@ uint8_t x_6; x_6 = !lean_is_exclusive(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_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; x_7 = lean_ctor_get(x_3, 0); x_8 = lean_ctor_get(x_3, 1); x_9 = lean_ctor_get(x_3, 2); x_10 = lean_ctor_get(x_3, 3); +x_11 = lean_ctor_get(x_4, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 1); +lean_inc(x_12); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_4); -x_11 = l_Lean_IR_LocalContext_addJP(x_4, x_7, x_8, x_9); +x_13 = l_Lean_IR_LocalContext_addJP(x_11, x_7, x_8, x_9); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); lean_inc(x_1); -x_12 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_10, x_11, x_5); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_ctor_get(x_13, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_13, 1); +x_15 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_10, x_14, x_5); +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_13); -x_17 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_9, x_4, x_14); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_9, x_4, x_17); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_17, 0); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_20, 0); +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_ctor_get(x_19, 1); -lean_dec(x_22); -lean_ctor_set(x_3, 3, x_15); -lean_ctor_set(x_3, 2, x_21); -lean_ctor_set(x_19, 1, x_16); -lean_ctor_set(x_19, 0, x_3); -return x_17; +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_22, 0); +x_25 = lean_ctor_get(x_22, 1); +lean_dec(x_25); +lean_ctor_set(x_3, 3, x_18); +lean_ctor_set(x_3, 2, x_24); +lean_ctor_set(x_22, 1, x_19); +lean_ctor_set(x_22, 0, x_3); +return x_20; } else { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -lean_dec(x_19); -lean_ctor_set(x_3, 3, x_15); -lean_ctor_set(x_3, 2, x_23); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_3); -lean_ctor_set(x_24, 1, x_16); -lean_ctor_set(x_17, 0, x_24); -return x_17; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_22, 0); +lean_inc(x_26); +lean_dec(x_22); +lean_ctor_set(x_3, 3, x_18); +lean_ctor_set(x_3, 2, x_26); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_3); +lean_ctor_set(x_27, 1, x_19); +lean_ctor_set(x_20, 0, x_27); +return x_20; } } 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; -x_25 = lean_ctor_get(x_17, 0); -x_26 = lean_ctor_get(x_17, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_17); -x_27 = lean_ctor_get(x_25, 0); -lean_inc(x_27); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_28 = x_25; +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_28 = lean_ctor_get(x_20, 0); +x_29 = lean_ctor_get(x_20, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_20); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +if (lean_is_exclusive(x_28)) { + lean_ctor_release(x_28, 0); + lean_ctor_release(x_28, 1); + x_31 = x_28; } else { - lean_dec_ref(x_25); - x_28 = lean_box(0); + lean_dec_ref(x_28); + x_31 = lean_box(0); } -lean_ctor_set(x_3, 3, x_15); -lean_ctor_set(x_3, 2, x_27); -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 3, x_18); +lean_ctor_set(x_3, 2, x_30); +if (lean_is_scalar(x_31)) { + x_32 = lean_alloc_ctor(0, 2, 0); } else { - x_29 = x_28; + x_32 = x_31; } -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_16); -x_30 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_26); -return x_30; +lean_ctor_set(x_32, 0, x_3); +lean_ctor_set(x_32, 1, x_19); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_29); +return x_33; } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_31 = lean_ctor_get(x_3, 0); -x_32 = lean_ctor_get(x_3, 1); -x_33 = lean_ctor_get(x_3, 2); -x_34 = lean_ctor_get(x_3, 3); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_34 = lean_ctor_get(x_3, 0); +x_35 = lean_ctor_get(x_3, 1); +x_36 = lean_ctor_get(x_3, 2); +x_37 = lean_ctor_get(x_3, 3); +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_3); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_4); -x_35 = l_Lean_IR_LocalContext_addJP(x_4, x_31, x_32, x_33); -lean_inc(x_1); -x_36 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_34, x_35, x_5); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +x_38 = lean_ctor_get(x_4, 0); lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 0); +x_39 = lean_ctor_get(x_4, 1); lean_inc(x_39); -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -x_41 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_33, x_4, x_38); -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +x_40 = l_Lean_IR_LocalContext_addJP(x_38, x_34, x_35, x_36); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_39); +lean_inc(x_1); +x_42 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_37, x_41, x_5); +x_43 = lean_ctor_get(x_42, 0); lean_inc(x_43); -if (lean_is_exclusive(x_41)) { - lean_ctor_release(x_41, 0); - lean_ctor_release(x_41, 1); - x_44 = x_41; +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = lean_ctor_get(x_43, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +x_47 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_36, x_4, x_44); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_50 = x_47; } else { - lean_dec_ref(x_41); - x_44 = lean_box(0); + lean_dec_ref(x_47); + x_50 = lean_box(0); } -x_45 = lean_ctor_get(x_42, 0); -lean_inc(x_45); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_46 = x_42; +x_51 = lean_ctor_get(x_48, 0); +lean_inc(x_51); +if (lean_is_exclusive(x_48)) { + lean_ctor_release(x_48, 0); + lean_ctor_release(x_48, 1); + x_52 = x_48; } else { - lean_dec_ref(x_42); - x_46 = lean_box(0); -} -x_47 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_47, 0, x_31); -lean_ctor_set(x_47, 1, x_32); -lean_ctor_set(x_47, 2, x_45); -lean_ctor_set(x_47, 3, x_39); -if (lean_is_scalar(x_46)) { - x_48 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_48); + x_52 = lean_box(0); +} +x_53 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_53, 0, x_34); +lean_ctor_set(x_53, 1, x_35); +lean_ctor_set(x_53, 2, x_51); +lean_ctor_set(x_53, 3, x_45); +if (lean_is_scalar(x_52)) { + x_54 = lean_alloc_ctor(0, 2, 0); } else { - x_48 = x_46; + x_54 = x_52; } -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_40); -if (lean_is_scalar(x_44)) { - x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_46); +if (lean_is_scalar(x_50)) { + x_55 = lean_alloc_ctor(0, 2, 0); } else { - x_49 = x_44; + x_55 = x_50; } -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_43); -return x_49; +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_49); +return x_55; } } case 10: { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_50 = lean_ctor_get(x_3, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_3, 1); -lean_inc(x_51); -x_52 = lean_ctor_get(x_3, 2); -lean_inc(x_52); -x_53 = lean_ctor_get(x_3, 3); -lean_inc(x_53); -lean_inc(x_4); +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; +x_56 = lean_ctor_get(x_3, 0); +lean_inc(x_56); +x_57 = lean_ctor_get(x_3, 1); +lean_inc(x_57); +x_58 = lean_ctor_get(x_3, 2); +lean_inc(x_58); +x_59 = lean_ctor_get(x_3, 3); +lean_inc(x_59); +x_60 = lean_ctor_get(x_4, 0); +lean_inc(x_60); lean_inc(x_3); -x_54 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_4, x_1); -x_55 = lean_unbox(x_54); -lean_dec(x_54); -if (x_55 == 0) +x_61 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_60, x_1); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) { -uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; -lean_dec(x_53); -lean_dec(x_52); -lean_dec(x_51); -lean_dec(x_50); +uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_59); +lean_dec(x_58); +lean_dec(x_57); +lean_dec(x_56); lean_dec(x_4); lean_dec(x_1); -x_56 = 0; -x_57 = lean_box(x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_3); -lean_ctor_set(x_58, 1, x_57); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_5); -return x_59; +x_63 = 0; +x_64 = lean_box(x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_3); +lean_ctor_set(x_65, 1, x_64); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_65); +lean_ctor_set(x_66, 1, x_5); +return x_66; } else { -uint8_t x_60; -x_60 = !lean_is_exclusive(x_3); -if (x_60 == 0) +uint8_t x_67; +x_67 = !lean_is_exclusive(x_3); +if (x_67 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; size_t x_66; size_t x_67; lean_object* x_68; uint8_t x_69; -x_61 = lean_ctor_get(x_3, 3); -lean_dec(x_61); -x_62 = lean_ctor_get(x_3, 2); -lean_dec(x_62); -x_63 = lean_ctor_get(x_3, 1); -lean_dec(x_63); -x_64 = lean_ctor_get(x_3, 0); -lean_dec(x_64); -x_65 = lean_array_get_size(x_53); -x_66 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_67 = 0; -x_68 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_66, x_67, x_53, x_4, x_5); -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_68, 0); -x_71 = lean_ctor_get(x_68, 1); -lean_ctor_set(x_3, 3, x_70); -x_72 = 1; -x_73 = lean_box(x_72); -lean_ctor_set(x_68, 1, x_73); -lean_ctor_set(x_68, 0, x_3); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_68); -lean_ctor_set(x_74, 1, x_71); -return x_74; -} -else -{ -lean_object* x_75; lean_object* x_76; uint8_t x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_75 = lean_ctor_get(x_68, 0); -x_76 = lean_ctor_get(x_68, 1); -lean_inc(x_76); -lean_inc(x_75); +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; uint8_t x_76; +x_68 = lean_ctor_get(x_3, 3); lean_dec(x_68); -lean_ctor_set(x_3, 3, x_75); -x_77 = 1; -x_78 = lean_box(x_77); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_3); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_76); -return x_80; -} -} -else -{ -lean_object* x_81; size_t x_82; size_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_69 = lean_ctor_get(x_3, 2); +lean_dec(x_69); +x_70 = lean_ctor_get(x_3, 1); +lean_dec(x_70); +x_71 = lean_ctor_get(x_3, 0); +lean_dec(x_71); +x_72 = lean_array_get_size(x_59); +x_73 = lean_usize_of_nat(x_72); +lean_dec(x_72); +x_74 = 0; +x_75 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_73, x_74, x_59, x_4, x_5); +x_76 = !lean_is_exclusive(x_75); +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_75, 0); +x_78 = lean_ctor_get(x_75, 1); +lean_ctor_set(x_3, 3, x_77); +x_79 = 1; +x_80 = lean_box(x_79); +lean_ctor_set(x_75, 1, x_80); +lean_ctor_set(x_75, 0, x_3); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_75); +lean_ctor_set(x_81, 1, x_78); +return x_81; +} +else +{ +lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +x_82 = lean_ctor_get(x_75, 0); +x_83 = lean_ctor_get(x_75, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_75); +lean_ctor_set(x_3, 3, x_82); +x_84 = 1; +x_85 = lean_box(x_84); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_3); +lean_ctor_set(x_86, 1, x_85); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_86); +lean_ctor_set(x_87, 1, x_83); +return x_87; +} +} +else +{ +lean_object* x_88; size_t x_89; size_t 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_dec(x_3); -x_81 = lean_array_get_size(x_53); -x_82 = lean_usize_of_nat(x_81); -lean_dec(x_81); -x_83 = 0; -x_84 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_82, x_83, x_53, x_4, x_5); -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -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; +x_88 = lean_array_get_size(x_59); +x_89 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_90 = 0; +x_91 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_89, x_90, x_59, x_4, x_5); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_94 = x_91; } else { - lean_dec_ref(x_84); - x_87 = lean_box(0); -} -x_88 = lean_alloc_ctor(10, 4, 0); -lean_ctor_set(x_88, 0, x_50); -lean_ctor_set(x_88, 1, x_51); -lean_ctor_set(x_88, 2, x_52); -lean_ctor_set(x_88, 3, x_85); -x_89 = 1; -x_90 = lean_box(x_89); -if (lean_is_scalar(x_87)) { - x_91 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_91); + x_94 = lean_box(0); +} +x_95 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_95, 0, x_56); +lean_ctor_set(x_95, 1, x_57); +lean_ctor_set(x_95, 2, x_58); +lean_ctor_set(x_95, 3, x_92); +x_96 = 1; +x_97 = lean_box(x_96); +if (lean_is_scalar(x_94)) { + x_98 = lean_alloc_ctor(0, 2, 0); } else { - x_91 = x_87; + x_98 = x_94; } -lean_ctor_set(x_91, 0, x_88); -lean_ctor_set(x_91, 1, x_90); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_86); -return x_92; +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_97); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_93); +return x_99; } } } default: { -uint8_t x_93; -x_93 = l_Lean_IR_FnBody_isTerminal(x_3); -if (x_93 == 0) +uint8_t x_100; +x_100 = l_Lean_IR_FnBody_isTerminal(x_3); +if (x_100 == 0) { -lean_object* x_94; lean_object* x_95; lean_object* x_96; uint8_t x_97; -x_94 = l_Lean_IR_FnBody_body(x_3); -x_95 = lean_box(13); +lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; +x_101 = l_Lean_IR_FnBody_body(x_3); +x_102 = lean_box(13); lean_inc(x_3); -x_96 = l_Lean_IR_FnBody_setBody(x_3, x_95); -x_97 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(x_96, x_1); -if (x_97 == 0) +x_103 = l_Lean_IR_FnBody_setBody(x_3, x_102); +x_104 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(x_103, x_1); +if (x_104 == 0) { -lean_object* x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; +lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; lean_dec(x_3); lean_inc(x_4); lean_inc(x_1); -x_98 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_94, x_4, x_5); -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_99, 1); -lean_inc(x_100); -x_101 = lean_unbox(x_100); -if (x_101 == 0) -{ -uint8_t x_102; -x_102 = !lean_is_exclusive(x_98); -if (x_102 == 0) -{ -lean_object* x_103; lean_object* x_104; uint8_t x_105; -x_103 = lean_ctor_get(x_98, 1); -x_104 = lean_ctor_get(x_98, 0); -lean_dec(x_104); -x_105 = !lean_is_exclusive(x_99); -if (x_105 == 0) -{ -lean_object* x_106; lean_object* x_107; uint8_t x_108; -x_106 = lean_ctor_get(x_99, 0); -x_107 = lean_ctor_get(x_99, 1); -lean_dec(x_107); -lean_inc(x_96); -x_108 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_96); +x_105 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_101, x_4, x_5); +x_106 = lean_ctor_get(x_105, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +x_108 = lean_unbox(x_107); if (x_108 == 0) { -lean_object* x_109; +uint8_t x_109; +x_109 = !lean_is_exclusive(x_105); +if (x_109 == 0) +{ +lean_object* x_110; lean_object* x_111; uint8_t x_112; +x_110 = lean_ctor_get(x_105, 1); +x_111 = lean_ctor_get(x_105, 0); +lean_dec(x_111); +x_112 = !lean_is_exclusive(x_106); +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; uint8_t x_115; +x_113 = lean_ctor_get(x_106, 0); +x_114 = lean_ctor_get(x_106, 1); +lean_dec(x_114); +lean_inc(x_103); +x_115 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); +if (x_115 == 0) +{ +lean_object* x_116; lean_dec(x_4); lean_dec(x_1); -x_109 = l_Lean_IR_FnBody_setBody(x_96, x_106); -lean_ctor_set(x_99, 0, x_109); -return x_98; +x_116 = l_Lean_IR_FnBody_setBody(x_103, x_113); +lean_ctor_set(x_106, 0, x_116); +return x_105; } else { -lean_object* x_110; uint8_t x_111; -lean_free_object(x_98); -lean_dec(x_100); -x_110 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_106, x_4, x_103); +lean_object* x_117; uint8_t x_118; +lean_free_object(x_105); +lean_dec(x_107); +x_117 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_113, x_4, x_110); lean_dec(x_4); -x_111 = !lean_is_exclusive(x_110); -if (x_111 == 0) -{ -lean_object* x_112; lean_object* x_113; lean_object* x_114; uint8_t x_115; lean_object* x_116; -x_112 = lean_ctor_get(x_110, 0); -x_113 = lean_ctor_get(x_110, 1); -x_114 = l_Lean_IR_FnBody_setBody(x_96, x_112); -x_115 = 1; -x_116 = lean_box(x_115); -lean_ctor_set(x_110, 1, x_116); -lean_ctor_set(x_110, 0, x_114); -lean_ctor_set(x_99, 1, x_113); -lean_ctor_set(x_99, 0, x_110); -return x_99; -} -else -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; lean_object* x_121; lean_object* x_122; -x_117 = lean_ctor_get(x_110, 0); -x_118 = lean_ctor_get(x_110, 1); -lean_inc(x_118); -lean_inc(x_117); -lean_dec(x_110); -x_119 = l_Lean_IR_FnBody_setBody(x_96, x_117); -x_120 = 1; -x_121 = lean_box(x_120); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_119); -lean_ctor_set(x_122, 1, x_121); -lean_ctor_set(x_99, 1, x_118); -lean_ctor_set(x_99, 0, x_122); -return x_99; -} -} -} -else -{ -lean_object* x_123; uint8_t x_124; -x_123 = lean_ctor_get(x_99, 0); -lean_inc(x_123); -lean_dec(x_99); -lean_inc(x_96); -x_124 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_96); -if (x_124 == 0) -{ -lean_object* x_125; lean_object* x_126; +x_118 = !lean_is_exclusive(x_117); +if (x_118 == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; +x_119 = lean_ctor_get(x_117, 0); +x_120 = lean_ctor_get(x_117, 1); +x_121 = l_Lean_IR_FnBody_setBody(x_103, x_119); +x_122 = 1; +x_123 = lean_box(x_122); +lean_ctor_set(x_117, 1, x_123); +lean_ctor_set(x_117, 0, x_121); +lean_ctor_set(x_106, 1, x_120); +lean_ctor_set(x_106, 0, x_117); +return x_106; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; +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 = l_Lean_IR_FnBody_setBody(x_103, x_124); +x_127 = 1; +x_128 = lean_box(x_127); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_126); +lean_ctor_set(x_129, 1, x_128); +lean_ctor_set(x_106, 1, x_125); +lean_ctor_set(x_106, 0, x_129); +return x_106; +} +} +} +else +{ +lean_object* x_130; uint8_t x_131; +x_130 = lean_ctor_get(x_106, 0); +lean_inc(x_130); +lean_dec(x_106); +lean_inc(x_103); +x_131 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; lean_dec(x_4); lean_dec(x_1); -x_125 = l_Lean_IR_FnBody_setBody(x_96, 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_100); -lean_ctor_set(x_98, 0, x_126); -return x_98; +x_132 = l_Lean_IR_FnBody_setBody(x_103, x_130); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_107); +lean_ctor_set(x_105, 0, x_133); +return x_105; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; -lean_free_object(x_98); -lean_dec(x_100); -x_127 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_123, x_4, x_103); +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_free_object(x_105); +lean_dec(x_107); +x_134 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_130, x_4, x_110); lean_dec(x_4); -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_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_137 = x_134; } else { - lean_dec_ref(x_127); - x_130 = lean_box(0); -} -x_131 = l_Lean_IR_FnBody_setBody(x_96, x_128); -x_132 = 1; -x_133 = lean_box(x_132); -if (lean_is_scalar(x_130)) { - x_134 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_134); + x_137 = lean_box(0); +} +x_138 = l_Lean_IR_FnBody_setBody(x_103, x_135); +x_139 = 1; +x_140 = lean_box(x_139); +if (lean_is_scalar(x_137)) { + x_141 = lean_alloc_ctor(0, 2, 0); } else { - x_134 = x_130; + x_141 = x_137; } -lean_ctor_set(x_134, 0, x_131); -lean_ctor_set(x_134, 1, x_133); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_129); -return x_135; +lean_ctor_set(x_141, 0, x_138); +lean_ctor_set(x_141, 1, x_140); +x_142 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_142, 0, x_141); +lean_ctor_set(x_142, 1, x_136); +return x_142; } } } else { -lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_136 = lean_ctor_get(x_98, 1); -lean_inc(x_136); -lean_dec(x_98); -x_137 = lean_ctor_get(x_99, 0); -lean_inc(x_137); -if (lean_is_exclusive(x_99)) { - lean_ctor_release(x_99, 0); - lean_ctor_release(x_99, 1); - x_138 = x_99; +lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; +x_143 = lean_ctor_get(x_105, 1); +lean_inc(x_143); +lean_dec(x_105); +x_144 = lean_ctor_get(x_106, 0); +lean_inc(x_144); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_145 = x_106; } else { - lean_dec_ref(x_99); - x_138 = lean_box(0); + lean_dec_ref(x_106); + x_145 = lean_box(0); } -lean_inc(x_96); -x_139 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_96); -if (x_139 == 0) +lean_inc(x_103); +x_146 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); +if (x_146 == 0) { -lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_dec(x_4); lean_dec(x_1); -x_140 = l_Lean_IR_FnBody_setBody(x_96, x_137); -if (lean_is_scalar(x_138)) { - x_141 = lean_alloc_ctor(0, 2, 0); +x_147 = l_Lean_IR_FnBody_setBody(x_103, x_144); +if (lean_is_scalar(x_145)) { + x_148 = lean_alloc_ctor(0, 2, 0); } else { - x_141 = x_138; + x_148 = x_145; } -lean_ctor_set(x_141, 0, x_140); -lean_ctor_set(x_141, 1, x_100); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_136); -return x_142; +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_148, 1, x_107); +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_148); +lean_ctor_set(x_149, 1, x_143); +return x_149; } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; -lean_dec(x_100); -x_143 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_137, x_4, x_136); +lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +lean_dec(x_107); +x_150 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_144, x_4, x_143); lean_dec(x_4); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_143)) { - lean_ctor_release(x_143, 0); - lean_ctor_release(x_143, 1); - x_146 = x_143; +x_151 = lean_ctor_get(x_150, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_150, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_150)) { + lean_ctor_release(x_150, 0); + lean_ctor_release(x_150, 1); + x_153 = x_150; } else { - lean_dec_ref(x_143); - x_146 = lean_box(0); -} -x_147 = l_Lean_IR_FnBody_setBody(x_96, x_144); -x_148 = 1; -x_149 = lean_box(x_148); -if (lean_is_scalar(x_146)) { - x_150 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_150); + x_153 = lean_box(0); +} +x_154 = l_Lean_IR_FnBody_setBody(x_103, x_151); +x_155 = 1; +x_156 = lean_box(x_155); +if (lean_is_scalar(x_153)) { + x_157 = lean_alloc_ctor(0, 2, 0); } else { - x_150 = x_146; + x_157 = x_153; } -lean_ctor_set(x_150, 0, x_147); -lean_ctor_set(x_150, 1, x_149); -if (lean_is_scalar(x_138)) { - x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_157, 0, x_154); +lean_ctor_set(x_157, 1, x_156); +if (lean_is_scalar(x_145)) { + x_158 = lean_alloc_ctor(0, 2, 0); } else { - x_151 = x_138; + x_158 = x_145; } -lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_145); -return x_151; +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_152); +return x_158; } } } else { -uint8_t x_152; +uint8_t x_159; lean_dec(x_4); lean_dec(x_1); -x_152 = !lean_is_exclusive(x_98); -if (x_152 == 0) -{ -lean_object* x_153; uint8_t x_154; -x_153 = lean_ctor_get(x_98, 0); -lean_dec(x_153); -x_154 = !lean_is_exclusive(x_99); -if (x_154 == 0) -{ -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_99, 0); -x_156 = lean_ctor_get(x_99, 1); -lean_dec(x_156); -x_157 = l_Lean_IR_FnBody_setBody(x_96, x_155); -lean_ctor_set(x_99, 0, x_157); -return x_98; -} -else -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_158 = lean_ctor_get(x_99, 0); -lean_inc(x_158); -lean_dec(x_99); -x_159 = l_Lean_IR_FnBody_setBody(x_96, x_158); -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_100); -lean_ctor_set(x_98, 0, x_160); -return x_98; -} -} -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; -x_161 = lean_ctor_get(x_98, 1); -lean_inc(x_161); -lean_dec(x_98); -x_162 = lean_ctor_get(x_99, 0); -lean_inc(x_162); -if (lean_is_exclusive(x_99)) { - lean_ctor_release(x_99, 0); - lean_ctor_release(x_99, 1); - x_163 = x_99; +x_159 = !lean_is_exclusive(x_105); +if (x_159 == 0) +{ +lean_object* x_160; uint8_t x_161; +x_160 = lean_ctor_get(x_105, 0); +lean_dec(x_160); +x_161 = !lean_is_exclusive(x_106); +if (x_161 == 0) +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_106, 0); +x_163 = lean_ctor_get(x_106, 1); +lean_dec(x_163); +x_164 = l_Lean_IR_FnBody_setBody(x_103, x_162); +lean_ctor_set(x_106, 0, x_164); +return x_105; +} +else +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_106, 0); +lean_inc(x_165); +lean_dec(x_106); +x_166 = l_Lean_IR_FnBody_setBody(x_103, x_165); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_166); +lean_ctor_set(x_167, 1, x_107); +lean_ctor_set(x_105, 0, x_167); +return x_105; +} +} +else +{ +lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_168 = lean_ctor_get(x_105, 1); +lean_inc(x_168); +lean_dec(x_105); +x_169 = lean_ctor_get(x_106, 0); +lean_inc(x_169); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_170 = x_106; } else { - lean_dec_ref(x_99); - x_163 = lean_box(0); + lean_dec_ref(x_106); + x_170 = lean_box(0); } -x_164 = l_Lean_IR_FnBody_setBody(x_96, x_162); -if (lean_is_scalar(x_163)) { - x_165 = lean_alloc_ctor(0, 2, 0); +x_171 = l_Lean_IR_FnBody_setBody(x_103, x_169); +if (lean_is_scalar(x_170)) { + x_172 = lean_alloc_ctor(0, 2, 0); } else { - x_165 = x_163; + x_172 = x_170; } -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_100); -x_166 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_166, 0, x_165); -lean_ctor_set(x_166, 1, x_161); -return x_166; +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_172, 1, x_107); +x_173 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_173, 0, x_172); +lean_ctor_set(x_173, 1, x_168); +return x_173; } } } else { -uint8_t x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_96); -lean_dec(x_94); +uint8_t x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_103); +lean_dec(x_101); lean_dec(x_4); lean_dec(x_1); -x_167 = 1; -x_168 = lean_box(x_167); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_3); -lean_ctor_set(x_169, 1, x_168); -x_170 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_5); -return x_170; +x_174 = 1; +x_175 = lean_box(x_174); +x_176 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_176, 0, x_3); +lean_ctor_set(x_176, 1, x_175); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_5); +return x_177; } } else { -lean_object* x_171; uint8_t x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; +x_178 = lean_ctor_get(x_4, 0); +lean_inc(x_178); +lean_dec(x_4); lean_inc(x_3); -x_171 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_4, x_1); +x_179 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_178, x_1); lean_dec(x_1); -x_172 = lean_unbox(x_171); -lean_dec(x_171); -x_173 = lean_box(x_172); -x_174 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_174, 0, x_3); -lean_ctor_set(x_174, 1, x_173); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_174); -lean_ctor_set(x_175, 1, x_5); -return x_175; +x_180 = lean_unbox(x_179); +lean_dec(x_179); +x_181 = lean_box(x_180); +x_182 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_182, 0, x_3); +lean_ctor_set(x_182, 1, x_181); +x_183 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_5); +return x_183; } } } @@ -1703,188 +1832,899 @@ lean_dec(x_2); return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -uint8_t x_7; -x_7 = lean_usize_dec_lt(x_3, x_2); +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_lt(x_4, x_6); +lean_dec(x_6); if (x_7 == 0) { -lean_object* x_8; -lean_dec(x_5); -lean_dec(x_1); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_4); -lean_ctor_set(x_8, 1, x_6); +uint8_t x_8; +lean_dec(x_4); +x_8 = 0; return x_8; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_array_uget(x_4, x_3); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_4, x_3, x_10); -if (lean_obj_tag(x_9) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_5); -x_15 = l_Lean_IR_ResetReuse_R(x_14, x_5, x_6); -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_16); -lean_inc(x_13); -lean_ctor_set(x_9, 1, x_16); -x_18 = l_Lean_IR_CtorInfo_isScalar(x_13); -if (x_18 == 0) -{ -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_9; uint8_t x_10; +x_9 = lean_array_fget(x_1, x_4); +x_10 = lean_nat_dec_eq(x_5, x_9); lean_dec(x_9); -lean_inc(x_5); -lean_inc(x_1); -x_19 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_13, x_16, x_5, x_17); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_13); -lean_ctor_set(x_22, 1, x_20); -x_23 = 1; -x_24 = lean_usize_add(x_3, x_23); -x_25 = lean_array_uset(x_11, x_3, x_22); -x_3 = x_24; -x_4 = x_25; -x_6 = x_21; +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_4, x_11); +lean_dec(x_4); +x_3 = lean_box(0); +x_4 = x_12; goto _start; } else { -size_t x_27; size_t x_28; lean_object* x_29; -lean_dec(x_16); -lean_dec(x_13); -x_27 = 1; -x_28 = lean_usize_add(x_3, x_27); -x_29 = lean_array_uset(x_11, x_3, x_9); -x_3 = x_28; -x_4 = x_29; -x_6 = x_17; -goto _start; +uint8_t x_14; +lean_dec(x_4); +x_14 = 1; +return x_14; } } -else +} +} +static size_t _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1() { +_start: { -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_31 = lean_ctor_get(x_9, 0); -x_32 = lean_ctor_get(x_9, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_9); -lean_inc(x_5); -x_33 = l_Lean_IR_ResetReuse_R(x_32, x_5, x_6); -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); -lean_inc(x_34); -lean_inc(x_31); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_31); -lean_ctor_set(x_36, 1, x_34); -x_37 = l_Lean_IR_CtorInfo_isScalar(x_31); -if (x_37 == 0) +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = 5; +x_3 = lean_usize_shift_left(x_1, x_2); +return x_3; +} +} +static size_t _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2() { +_start: { -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_dec(x_36); -lean_inc(x_5); -lean_inc(x_1); -x_38 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_31, x_34, x_5, x_35); -x_39 = lean_ctor_get(x_38, 0); +size_t x_1; size_t x_2; size_t x_3; +x_1 = 1; +x_2 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1; +x_3 = lean_usize_sub(x_2, x_1); +return x_3; +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = 5; +x_6 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2; +x_7 = lean_usize_land(x_2, x_6); +x_8 = lean_usize_to_nat(x_7); +x_9 = lean_box(2); +x_10 = lean_array_get(x_9, x_4, x_8); +lean_dec(x_8); +lean_dec(x_4); +switch (lean_obj_tag(x_10)) { +case 0: +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = lean_nat_dec_eq(x_3, x_11); +lean_dec(x_11); +return x_12; +} +case 1: +{ +lean_object* x_13; size_t x_14; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_usize_shift_right(x_2, x_5); +x_1 = x_13; +x_2 = x_14; +goto _start; +} +default: +{ +uint8_t x_16; +x_16 = 0; +return x_16; +} +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +lean_dec(x_1); +x_19 = lean_unsigned_to_nat(0u); +x_20 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3(x_17, x_18, lean_box(0), x_19, x_3); +lean_dec(x_18); +lean_dec(x_17); +return x_20; +} +} +} +LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; uint64_t x_4; size_t x_5; uint8_t x_6; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +lean_dec(x_1); +x_4 = lean_uint64_of_nat(x_2); +x_5 = lean_uint64_to_usize(x_4); +x_6 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2(x_3, x_5, x_2); +lean_dec(x_2); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_dec(x_5); +return x_6; +} +else +{ +lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = lean_array_fget(x_2, x_5); +x_10 = lean_array_fget(x_3, x_5); +x_11 = lean_uint64_of_nat(x_9); +x_12 = lean_uint64_to_usize(x_11); +x_13 = 1; +x_14 = lean_usize_sub(x_1, x_13); +x_15 = 5; +x_16 = lean_usize_mul(x_15, x_14); +x_17 = lean_usize_shift_right(x_12, x_16); +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_add(x_5, x_18); +lean_dec(x_5); +x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_6, x_17, x_1, x_9, x_10); +x_4 = lean_box(0); +x_5 = x_19; +x_6 = x_20; +goto _start; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_IR_ResetReuse_R___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_1, 1); +lean_inc(x_6); +x_7 = lean_array_get_size(x_5); +x_8 = lean_nat_dec_lt(x_2, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_2); +x_9 = !lean_is_exclusive(x_1); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_1, 1); +lean_dec(x_10); +x_11 = lean_ctor_get(x_1, 0); +lean_dec(x_11); +x_12 = lean_array_push(x_5, x_3); +x_13 = lean_array_push(x_6, x_4); +lean_ctor_set(x_1, 1, x_13); +lean_ctor_set(x_1, 0, x_12); +return x_1; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_14 = lean_array_push(x_5, x_3); +x_15 = lean_array_push(x_6, x_4); +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set(x_16, 1, x_15); +return x_16; +} +} +else +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_array_fget(x_5, x_2); +x_18 = lean_nat_dec_eq(x_3, x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_6); +lean_dec(x_5); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_2, x_19); +lean_dec(x_2); +x_2 = x_20; +goto _start; +} +else +{ +uint8_t x_22; +x_22 = !lean_is_exclusive(x_1); +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_1, 1); +lean_dec(x_23); +x_24 = lean_ctor_get(x_1, 0); +lean_dec(x_24); +x_25 = lean_array_fset(x_5, x_2, x_3); +x_26 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +lean_ctor_set(x_1, 1, x_26); +lean_ctor_set(x_1, 0, x_25); +return x_1; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_1); +x_27 = lean_array_fset(x_5, x_2, x_3); +x_28 = lean_array_fset(x_6, x_2, x_4); +lean_dec(x_2); +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; +} +} +} +} +} +static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +uint8_t x_6; +x_6 = !lean_is_exclusive(x_1); +if (x_6 == 0) +{ +lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_7 = lean_ctor_get(x_1, 0); +x_8 = 1; +x_9 = 5; +x_10 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2; +x_11 = lean_usize_land(x_2, x_10); +x_12 = lean_usize_to_nat(x_11); +x_13 = lean_array_get_size(x_7); +x_14 = lean_nat_dec_lt(x_12, x_13); +lean_dec(x_13); +if (x_14 == 0) +{ +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +return x_1; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_12); +x_16 = lean_box(0); +x_17 = lean_array_fset(x_7, x_12, x_16); +switch (lean_obj_tag(x_15)) { +case 0: +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +x_21 = lean_nat_dec_eq(x_4, x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +lean_free_object(x_15); +x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +x_24 = lean_array_fset(x_17, x_12, x_23); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_24); +return x_1; +} +else +{ +lean_object* x_25; +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_25 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_25); +return x_1; +} +} +else +{ +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_15); +x_28 = lean_nat_dec_eq(x_4, x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, 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; +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_27); +lean_dec(x_26); +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_4); +lean_ctor_set(x_32, 1, x_5); +x_33 = lean_array_fset(x_17, x_12, x_32); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_33); +return x_1; +} +} +} +case 1: +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) +{ +lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_usize_shift_right(x_2, x_9); +x_37 = lean_usize_add(x_3, x_8); +x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_35, x_36, x_37, x_4, x_5); +lean_ctor_set(x_15, 0, x_38); +x_39 = lean_array_fset(x_17, x_12, x_15); +lean_dec(x_12); +lean_ctor_set(x_1, 0, x_39); +return x_1; +} +else +{ +lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_40 = lean_ctor_get(x_15, 0); +lean_inc(x_40); +lean_dec(x_15); +x_41 = lean_usize_shift_right(x_2, x_9); +x_42 = lean_usize_add(x_3, x_8); +x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_40, x_41, x_42, 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_dec(x_12); +lean_ctor_set(x_1, 0, x_45); +return x_1; +} +} +default: +{ +lean_object* x_46; lean_object* x_47; +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_dec(x_12); +lean_ctor_set(x_1, 0, x_47); +return x_1; +} +} +} +} +else +{ +lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = 1; +x_50 = 5; +x_51 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2; +x_52 = lean_usize_land(x_2, x_51); +x_53 = lean_usize_to_nat(x_52); +x_54 = lean_array_get_size(x_48); +x_55 = lean_nat_dec_lt(x_53, x_54); +lean_dec(x_54); +if (x_55 == 0) +{ +lean_object* x_56; +lean_dec(x_53); +lean_dec(x_5); +lean_dec(x_4); +x_56 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_56, 0, x_48); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_array_fget(x_48, x_53); +x_58 = lean_box(0); +x_59 = lean_array_fset(x_48, x_53, x_58); +switch (lean_obj_tag(x_57)) { +case 0: +{ +lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_62 = x_57; +} else { + lean_dec_ref(x_57); + x_62 = lean_box(0); +} +x_63 = lean_nat_dec_eq(x_4, x_60); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_62); +x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); +x_65 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_65, 0, x_64); +x_66 = lean_array_fset(x_59, x_53, x_65); +lean_dec(x_53); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_66); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_dec(x_61); +lean_dec(x_60); +if (lean_is_scalar(x_62)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_62; +} +lean_ctor_set(x_68, 0, x_4); +lean_ctor_set(x_68, 1, x_5); +x_69 = lean_array_fset(x_59, x_53, x_68); +lean_dec(x_53); +x_70 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_70, 0, x_69); +return x_70; +} +} +case 1: +{ +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; +x_71 = lean_ctor_get(x_57, 0); +lean_inc(x_71); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + x_72 = x_57; +} else { + lean_dec_ref(x_57); + x_72 = lean_box(0); +} +x_73 = lean_usize_shift_right(x_2, x_50); +x_74 = lean_usize_add(x_3, x_49); +x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_71, x_73, x_74, x_4, x_5); +if (lean_is_scalar(x_72)) { + x_76 = lean_alloc_ctor(1, 1, 0); +} else { + x_76 = x_72; +} +lean_ctor_set(x_76, 0, x_75); +x_77 = lean_array_fset(x_59, x_53, x_76); +lean_dec(x_53); +x_78 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_78, 0, x_77); +return x_78; +} +default: +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_4); +lean_ctor_set(x_79, 1, x_5); +x_80 = lean_array_fset(x_59, x_53, x_79); +lean_dec(x_53); +x_81 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_81, 0, x_80); +return x_81; +} +} +} +} +} +else +{ +uint8_t x_82; +x_82 = !lean_is_exclusive(x_1); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; +x_83 = lean_unsigned_to_nat(0u); +x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_IR_ResetReuse_R___spec__7(x_1, x_83, x_4, x_5); +x_85 = 7; +x_86 = lean_usize_dec_le(x_85, x_3); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; uint8_t x_89; +x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); +x_88 = lean_unsigned_to_nat(4u); +x_89 = lean_nat_dec_lt(x_87, x_88); +lean_dec(x_87); +if (x_89 == 0) +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_84, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_84, 1); +lean_inc(x_91); +lean_dec(x_84); +x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1; +x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(x_3, x_90, x_91, lean_box(0), x_83, x_92); +lean_dec(x_91); +lean_dec(x_90); +return x_93; +} +else +{ +return x_84; +} +} +else +{ +return x_84; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; +x_94 = lean_ctor_get(x_1, 0); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_1); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +x_97 = lean_unsigned_to_nat(0u); +x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_IR_ResetReuse_R___spec__7(x_96, x_97, x_4, x_5); +x_99 = 7; +x_100 = lean_usize_dec_le(x_99, x_3); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; uint8_t x_103; +x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); +x_102 = lean_unsigned_to_nat(4u); +x_103 = lean_nat_dec_lt(x_101, x_102); +lean_dec(x_101); +if (x_103 == 0) +{ +lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_104 = lean_ctor_get(x_98, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1; +x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(x_3, x_104, x_105, lean_box(0), x_97, x_106); +lean_dec(x_105); +lean_dec(x_104); +return x_107; +} +else +{ +return x_98; +} +} +else +{ +return x_98; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = lean_uint64_of_nat(x_2); +x_8 = lean_uint64_to_usize(x_7); +x_9 = 1; +x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_5, x_8, x_9, x_2, x_3); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_add(x_6, x_11); +lean_dec(x_6); +lean_ctor_set(x_1, 1, x_12); +lean_ctor_set(x_1, 0, x_10); +return x_1; +} +else +{ +lean_object* x_13; lean_object* x_14; uint64_t 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; +x_13 = lean_ctor_get(x_1, 0); +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_1); +x_15 = lean_uint64_of_nat(x_2); +x_16 = lean_uint64_to_usize(x_15); +x_17 = 1; +x_18 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_13, x_16, x_17, x_2, x_3); +x_19 = lean_unsigned_to_nat(1u); +x_20 = lean_nat_add(x_14, x_19); +lean_dec(x_14); +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; +} +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(lean_object* x_1, uint8_t x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = lean_usize_dec_lt(x_4, x_3); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_6); +lean_dec(x_1); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_5); +lean_ctor_set(x_9, 1, x_7); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_uget(x_5, x_4); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_array_uset(x_5, x_4, x_11); +if (lean_obj_tag(x_10) == 0) +{ +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; uint8_t x_19; +x_14 = lean_ctor_get(x_10, 0); +x_15 = lean_ctor_get(x_10, 1); +lean_inc(x_6); +x_16 = l_Lean_IR_ResetReuse_R(x_15, x_6, x_7); +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_17); +lean_inc(x_14); +lean_ctor_set(x_10, 1, x_17); +x_19 = l_Lean_IR_CtorInfo_isScalar(x_14); +if (x_19 == 0) +{ +if (x_2 == 0) +{ +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_dec(x_10); +lean_inc(x_6); +lean_inc(x_1); +x_20 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_14, x_17, x_6, 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 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_14); +lean_ctor_set(x_23, 1, x_21); +x_24 = 1; +x_25 = lean_usize_add(x_4, x_24); +x_26 = lean_array_uset(x_12, x_4, x_23); +x_4 = x_25; +x_5 = x_26; +x_7 = x_22; +goto _start; +} +else +{ +size_t x_28; size_t x_29; lean_object* x_30; +lean_dec(x_17); +lean_dec(x_14); +x_28 = 1; +x_29 = lean_usize_add(x_4, x_28); +x_30 = lean_array_uset(x_12, x_4, x_10); +x_4 = x_29; +x_5 = x_30; +x_7 = x_18; +goto _start; +} +} +else +{ +size_t x_32; size_t x_33; lean_object* x_34; +lean_dec(x_17); +lean_dec(x_14); +x_32 = 1; +x_33 = lean_usize_add(x_4, x_32); +x_34 = lean_array_uset(x_12, x_4, x_10); +x_4 = x_33; +x_5 = x_34; +x_7 = x_18; +goto _start; +} +} +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; uint8_t x_42; +x_36 = lean_ctor_get(x_10, 0); +x_37 = lean_ctor_get(x_10, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_10); +lean_inc(x_6); +x_38 = l_Lean_IR_ResetReuse_R(x_37, x_6, x_7); +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); +lean_inc(x_39); +lean_inc(x_36); x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_31); +lean_ctor_set(x_41, 0, x_36); lean_ctor_set(x_41, 1, x_39); -x_42 = 1; -x_43 = lean_usize_add(x_3, x_42); -x_44 = lean_array_uset(x_11, x_3, x_41); -x_3 = x_43; -x_4 = x_44; -x_6 = x_40; +x_42 = l_Lean_IR_CtorInfo_isScalar(x_36); +if (x_42 == 0) +{ +if (x_2 == 0) +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; size_t x_48; lean_object* x_49; +lean_dec(x_41); +lean_inc(x_6); +lean_inc(x_1); +x_43 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(x_1, x_36, x_39, x_6, x_40); +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_alloc_ctor(0, 2, 0); +lean_ctor_set(x_46, 0, x_36); +lean_ctor_set(x_46, 1, x_44); +x_47 = 1; +x_48 = lean_usize_add(x_4, x_47); +x_49 = lean_array_uset(x_12, x_4, x_46); +x_4 = x_48; +x_5 = x_49; +x_7 = x_45; goto _start; } else { -size_t x_46; size_t x_47; lean_object* x_48; -lean_dec(x_34); -lean_dec(x_31); -x_46 = 1; -x_47 = lean_usize_add(x_3, x_46); -x_48 = lean_array_uset(x_11, x_3, x_36); -x_3 = x_47; -x_4 = x_48; -x_6 = x_35; +size_t x_51; size_t x_52; lean_object* x_53; +lean_dec(x_39); +lean_dec(x_36); +x_51 = 1; +x_52 = lean_usize_add(x_4, x_51); +x_53 = lean_array_uset(x_12, x_4, x_41); +x_4 = x_52; +x_5 = x_53; +x_7 = x_40; goto _start; } } -} else { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_9); -if (x_50 == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; -x_51 = lean_ctor_get(x_9, 0); -lean_inc(x_5); -x_52 = l_Lean_IR_ResetReuse_R(x_51, x_5, x_6); -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_ctor_set(x_9, 0, x_53); +size_t x_55; size_t x_56; lean_object* x_57; +lean_dec(x_39); +lean_dec(x_36); x_55 = 1; -x_56 = lean_usize_add(x_3, x_55); -x_57 = lean_array_uset(x_11, x_3, x_9); -x_3 = x_56; -x_4 = x_57; -x_6 = x_54; +x_56 = lean_usize_add(x_4, x_55); +x_57 = lean_array_uset(x_12, x_4, x_41); +x_4 = x_56; +x_5 = x_57; +x_7 = x_40; goto _start; } +} +} else { -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; -x_59 = lean_ctor_get(x_9, 0); -lean_inc(x_59); -lean_dec(x_9); -lean_inc(x_5); -x_60 = l_Lean_IR_ResetReuse_R(x_59, x_5, x_6); -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); +uint8_t x_59; +x_59 = !lean_is_exclusive(x_10); +if (x_59 == 0) +{ +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; +x_60 = lean_ctor_get(x_10, 0); +lean_inc(x_6); +x_61 = l_Lean_IR_ResetReuse_R(x_60, x_6, x_7); +x_62 = lean_ctor_get(x_61, 0); lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_63, 0, x_61); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +lean_ctor_set(x_10, 0, x_62); x_64 = 1; -x_65 = lean_usize_add(x_3, x_64); -x_66 = lean_array_uset(x_11, x_3, x_63); -x_3 = x_65; -x_4 = x_66; -x_6 = x_62; +x_65 = lean_usize_add(x_4, x_64); +x_66 = lean_array_uset(x_12, x_4, x_10); +x_4 = x_65; +x_5 = x_66; +x_7 = x_63; +goto _start; +} +else +{ +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; +x_68 = lean_ctor_get(x_10, 0); +lean_inc(x_68); +lean_dec(x_10); +lean_inc(x_6); +x_69 = l_Lean_IR_ResetReuse_R(x_68, x_6, x_7); +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_72, 0, x_70); +x_73 = 1; +x_74 = lean_usize_add(x_4, x_73); +x_75 = lean_array_uset(x_12, x_4, x_72); +x_4 = x_74; +x_5 = x_75; +x_7 = x_71; goto _start; } } @@ -1901,7 +2741,7 @@ uint8_t x_4; x_4 = !lean_is_exclusive(x_1); if (x_4 == 0) { -lean_object* x_5; lean_object* x_6; lean_object* x_7; 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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_5 = lean_ctor_get(x_1, 0); x_6 = lean_ctor_get(x_1, 1); x_7 = lean_ctor_get(x_1, 2); @@ -1913,230 +2753,451 @@ 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_2); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_2, 0); lean_inc(x_10); lean_inc(x_6); lean_inc(x_5); -x_12 = l_Lean_IR_LocalContext_addJP(x_2, x_5, x_6, x_10); -x_13 = l_Lean_IR_ResetReuse_R(x_8, x_12, x_11); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_ctor_set(x_1, 3, x_15); +x_14 = l_Lean_IR_LocalContext_addJP(x_13, x_5, x_6, x_10); +lean_ctor_set(x_2, 0, x_14); +x_15 = l_Lean_IR_ResetReuse_R(x_8, x_2, x_11); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_ctor_set(x_1, 3, x_17); lean_ctor_set(x_1, 2, x_10); -lean_ctor_set(x_13, 0, x_1); -return x_13; +lean_ctor_set(x_15, 0, x_1); +return x_15; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_13, 0); -x_17 = lean_ctor_get(x_13, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_13); -lean_ctor_set(x_1, 3, x_16); +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_15, 0); +x_19 = lean_ctor_get(x_15, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_15); +lean_ctor_set(x_1, 3, x_18); lean_ctor_set(x_1, 2, x_10); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_17); -return x_18; +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_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_19 = lean_ctor_get(x_1, 0); -x_20 = lean_ctor_get(x_1, 1); -x_21 = lean_ctor_get(x_1, 2); -x_22 = lean_ctor_get(x_1, 3); +lean_object* x_21; lean_object* x_22; lean_object* x_23; 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_21 = lean_ctor_get(x_2, 0); +x_22 = lean_ctor_get(x_2, 1); lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); +lean_dec(x_2); +lean_inc(x_10); +lean_inc(x_6); +lean_inc(x_5); +x_23 = l_Lean_IR_LocalContext_addJP(x_21, x_5, x_6, x_10); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_22); +x_25 = l_Lean_IR_ResetReuse_R(x_8, x_24, x_11); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_25)) { + lean_ctor_release(x_25, 0); + lean_ctor_release(x_25, 1); + x_28 = x_25; +} else { + lean_dec_ref(x_25); + x_28 = lean_box(0); +} +lean_ctor_set(x_1, 3, x_26); +lean_ctor_set(x_1, 2, x_10); +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; +} +lean_ctor_set(x_29, 0, x_1); +lean_ctor_set(x_29, 1, x_27); +return x_29; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_30 = lean_ctor_get(x_1, 0); +x_31 = lean_ctor_get(x_1, 1); +x_32 = lean_ctor_get(x_1, 2); +x_33 = lean_ctor_get(x_1, 3); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); lean_dec(x_1); lean_inc(x_2); -x_23 = l_Lean_IR_ResetReuse_R(x_21, x_2, x_3); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -lean_inc(x_24); -lean_inc(x_20); -lean_inc(x_19); -x_26 = l_Lean_IR_LocalContext_addJP(x_2, x_19, x_20, x_24); -x_27 = l_Lean_IR_ResetReuse_R(x_22, x_26, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_30 = x_27; +x_34 = l_Lean_IR_ResetReuse_R(x_32, x_2, x_3); +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_2, 0); +lean_inc(x_37); +x_38 = lean_ctor_get(x_2, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + x_39 = x_2; } else { - lean_dec_ref(x_27); - x_30 = lean_box(0); -} -x_31 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -lean_ctor_set(x_31, 2, x_24); -lean_ctor_set(x_31, 3, x_28); -if (lean_is_scalar(x_30)) { - x_32 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_2); + x_39 = lean_box(0); +} +lean_inc(x_35); +lean_inc(x_31); +lean_inc(x_30); +x_40 = l_Lean_IR_LocalContext_addJP(x_37, x_30, x_31, x_35); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); } else { - x_32 = x_30; + x_41 = x_39; } -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_29); -return x_32; +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +x_42 = l_Lean_IR_ResetReuse_R(x_33, x_41, x_36); +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +x_44 = lean_ctor_get(x_42, 1); +lean_inc(x_44); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_45 = x_42; +} else { + lean_dec_ref(x_42); + x_45 = lean_box(0); +} +x_46 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_46, 0, x_30); +lean_ctor_set(x_46, 1, x_31); +lean_ctor_set(x_46, 2, x_35); +lean_ctor_set(x_46, 3, x_43); +if (lean_is_scalar(x_45)) { + x_47 = lean_alloc_ctor(0, 2, 0); +} else { + x_47 = x_45; +} +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_44); +return x_47; } } case 10: { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_1); -if (x_33 == 0) +uint8_t x_48; +x_48 = !lean_is_exclusive(x_1); +if (x_48 == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; lean_object* x_39; uint8_t x_40; -x_34 = lean_ctor_get(x_1, 1); -x_35 = lean_ctor_get(x_1, 3); -x_36 = lean_array_get_size(x_35); -x_37 = lean_usize_of_nat(x_36); -lean_dec(x_36); -x_38 = 0; -lean_inc(x_34); -x_39 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1(x_34, x_37, x_38, x_35, x_2, x_3); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +uint8_t x_49; +x_49 = !lean_is_exclusive(x_2); +if (x_49 == 0) { -lean_object* x_41; -x_41 = lean_ctor_get(x_39, 0); -lean_ctor_set(x_1, 3, x_41); -lean_ctor_set(x_39, 0, x_1); -return x_39; +lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_50 = lean_ctor_get(x_1, 1); +x_51 = lean_ctor_get(x_1, 3); +x_52 = lean_ctor_get(x_2, 1); +lean_inc(x_50); +lean_inc(x_52); +x_53 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_52, x_50); +x_54 = lean_array_get_size(x_51); +x_55 = lean_usize_of_nat(x_54); +lean_dec(x_54); +x_56 = 0; +x_57 = lean_box(0); +lean_inc(x_50); +x_58 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_52, x_50, x_57); +lean_ctor_set(x_2, 1, x_58); +lean_inc(x_50); +x_59 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_50, x_53, x_55, x_56, x_51, x_2, x_3); +x_60 = !lean_is_exclusive(x_59); +if (x_60 == 0) +{ +lean_object* x_61; +x_61 = lean_ctor_get(x_59, 0); +lean_ctor_set(x_1, 3, x_61); +lean_ctor_set(x_59, 0, x_1); +return x_59; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_39, 0); -x_43 = lean_ctor_get(x_39, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_39); -lean_ctor_set(x_1, 3, x_42); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_1); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_59, 0); +x_63 = lean_ctor_get(x_59, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_59); +lean_ctor_set(x_1, 3, x_62); +x_64 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_64, 0, x_1); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; size_t x_50; size_t x_51; 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_45 = lean_ctor_get(x_1, 0); -x_46 = lean_ctor_get(x_1, 1); -x_47 = lean_ctor_get(x_1, 2); -x_48 = lean_ctor_get(x_1, 3); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; size_t x_71; size_t 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_65 = lean_ctor_get(x_1, 1); +x_66 = lean_ctor_get(x_1, 3); +x_67 = lean_ctor_get(x_2, 0); +x_68 = lean_ctor_get(x_2, 1); +lean_inc(x_68); +lean_inc(x_67); +lean_dec(x_2); +lean_inc(x_65); +lean_inc(x_68); +x_69 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_68, x_65); +x_70 = lean_array_get_size(x_66); +x_71 = lean_usize_of_nat(x_70); +lean_dec(x_70); +x_72 = 0; +x_73 = lean_box(0); +lean_inc(x_65); +x_74 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_68, x_65, x_73); +x_75 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_75, 0, x_67); +lean_ctor_set(x_75, 1, x_74); +lean_inc(x_65); +x_76 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_65, x_69, x_71, x_72, x_66, x_75, x_3); +x_77 = lean_ctor_get(x_76, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_76, 1); +lean_inc(x_78); +if (lean_is_exclusive(x_76)) { + lean_ctor_release(x_76, 0); + lean_ctor_release(x_76, 1); + x_79 = x_76; +} else { + lean_dec_ref(x_76); + x_79 = lean_box(0); +} +lean_ctor_set(x_1, 3, x_77); +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(0, 2, 0); +} else { + x_80 = x_79; +} +lean_ctor_set(x_80, 0, x_1); +lean_ctor_set(x_80, 1, x_78); +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; uint8_t x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_81 = lean_ctor_get(x_1, 0); +x_82 = lean_ctor_get(x_1, 1); +x_83 = lean_ctor_get(x_1, 2); +x_84 = lean_ctor_get(x_1, 3); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); lean_dec(x_1); -x_49 = lean_array_get_size(x_48); -x_50 = lean_usize_of_nat(x_49); -lean_dec(x_49); -x_51 = 0; -lean_inc(x_46); -x_52 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1(x_46, x_50, x_51, x_48, x_2, x_3); -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; +x_85 = lean_ctor_get(x_2, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_2, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_2)) { + lean_ctor_release(x_2, 0); + lean_ctor_release(x_2, 1); + x_87 = x_2; +} else { + lean_dec_ref(x_2); + x_87 = lean_box(0); +} +lean_inc(x_82); +lean_inc(x_86); +x_88 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_86, x_82); +x_89 = lean_array_get_size(x_84); +x_90 = lean_usize_of_nat(x_89); +lean_dec(x_89); +x_91 = 0; +x_92 = lean_box(0); +lean_inc(x_82); +x_93 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_86, x_82, x_92); +if (lean_is_scalar(x_87)) { + x_94 = lean_alloc_ctor(0, 2, 0); +} else { + x_94 = x_87; +} +lean_ctor_set(x_94, 0, x_85); +lean_ctor_set(x_94, 1, x_93); +lean_inc(x_82); +x_95 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_82, x_88, x_90, x_91, x_84, x_94, x_3); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_98 = x_95; } else { - lean_dec_ref(x_52); - x_55 = lean_box(0); -} -x_56 = lean_alloc_ctor(10, 4, 0); -lean_ctor_set(x_56, 0, x_45); -lean_ctor_set(x_56, 1, x_46); -lean_ctor_set(x_56, 2, x_47); -lean_ctor_set(x_56, 3, x_53); -if (lean_is_scalar(x_55)) { - x_57 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_95); + x_98 = lean_box(0); +} +x_99 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_99, 0, x_81); +lean_ctor_set(x_99, 1, x_82); +lean_ctor_set(x_99, 2, x_83); +lean_ctor_set(x_99, 3, x_96); +if (lean_is_scalar(x_98)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_55; + x_100 = x_98; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_97); +return x_100; } } default: { -uint8_t x_58; -x_58 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_58 == 0) +uint8_t x_101; +x_101 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_101 == 0) { -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_59 = l_Lean_IR_FnBody_body(x_1); -x_60 = lean_box(13); -x_61 = l_Lean_IR_FnBody_setBody(x_1, x_60); -x_62 = l_Lean_IR_ResetReuse_R(x_59, x_2, x_3); -x_63 = !lean_is_exclusive(x_62); -if (x_63 == 0) +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_102 = l_Lean_IR_FnBody_body(x_1); +x_103 = lean_box(13); +x_104 = l_Lean_IR_FnBody_setBody(x_1, x_103); +x_105 = l_Lean_IR_ResetReuse_R(x_102, x_2, x_3); +x_106 = !lean_is_exclusive(x_105); +if (x_106 == 0) { -lean_object* x_64; lean_object* x_65; -x_64 = lean_ctor_get(x_62, 0); -x_65 = l_Lean_IR_FnBody_setBody(x_61, x_64); -lean_ctor_set(x_62, 0, x_65); -return x_62; +lean_object* x_107; lean_object* x_108; +x_107 = lean_ctor_get(x_105, 0); +x_108 = l_Lean_IR_FnBody_setBody(x_104, x_107); +lean_ctor_set(x_105, 0, x_108); +return x_105; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_62, 0); -x_67 = lean_ctor_get(x_62, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_62); -x_68 = l_Lean_IR_FnBody_setBody(x_61, x_66); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_67); -return x_69; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_109 = lean_ctor_get(x_105, 0); +x_110 = lean_ctor_get(x_105, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_105); +x_111 = l_Lean_IR_FnBody_setBody(x_104, x_109); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_110); +return x_112; } } else { -lean_object* x_70; +lean_object* x_113; lean_dec(x_2); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_1); -lean_ctor_set(x_70, 1, x_3); -return x_70; +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_1); +lean_ctor_set(x_113, 1, x_3); +return x_113; } } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___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_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___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_7; size_t x_8; lean_object* x_9; -x_7 = lean_unbox_usize(x_2); +uint8_t x_6; lean_object* x_7; +x_6 = l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); lean_dec(x_2); -x_8 = lean_unbox_usize(x_3); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +size_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_5 = l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2(x_1, x_4, x_3); lean_dec(x_3); -x_9 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__1(x_1, x_7, x_8, x_4, x_5, x_6); -return x_9; +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___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: +{ +size_t x_7; lean_object* x_8; +x_7 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(x_7, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(x_1, x_6, x_7, x_4, x_5); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___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: +{ +uint8_t x_8; size_t x_9; size_t x_10; lean_object* x_11; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_1, x_8, x_9, x_10, x_5, x_6, x_7); +return x_11; +} +} +static lean_object* _init_l_Lean_IR_Decl_insertResetReuse___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object* x_1) { @@ -2153,7 +3214,7 @@ x_4 = l_Lean_IR_MaxIndex_collectDecl(x_1, x_3); x_5 = lean_unsigned_to_nat(1u); x_6 = lean_nat_add(x_4, x_5); lean_dec(x_4); -x_7 = lean_box(0); +x_7 = l_Lean_IR_Decl_insertResetReuse___closed__1; x_8 = l_Lean_IR_ResetReuse_R(x_2, x_7, x_6); x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); @@ -2184,6 +3245,28 @@ lean_dec_ref(res); res = initialize_Lean_Compiler_IR_Format(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_IR_ResetReuse_Context_lctx___default = _init_l_Lean_IR_ResetReuse_Context_lctx___default(); +lean_mark_persistent(l_Lean_IR_ResetReuse_Context_lctx___default); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1); +l_Lean_IR_ResetReuse_Context_casesVars___default = _init_l_Lean_IR_ResetReuse_Context_casesVars___default(); +lean_mark_persistent(l_Lean_IR_ResetReuse_Context_casesVars___default); +l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1(); +l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2(); +l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1); +l_Lean_IR_Decl_insertResetReuse___closed__1 = _init_l_Lean_IR_Decl_insertResetReuse___closed__1(); +lean_mark_persistent(l_Lean_IR_Decl_insertResetReuse___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/InitAttr.c b/stage0/stdlib/Lean/Compiler/InitAttr.c index 41fbb3078c8e..dd5a2a09bcc8 100644 --- a/stage0/stdlib/Lean/Compiler/InitAttr.c +++ b/stage0/stdlib/Lean/Compiler/InitAttr.c @@ -31,6 +31,7 @@ lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___closed__2; lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerInitAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_InitAttr_0__Lean_getIOTypeArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_setBuiltinInitAttr(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_declareBuiltin___closed__13; @@ -47,6 +48,7 @@ uint8_t l_Lean_Name_isAnonymous(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__10; static lean_object* l_Lean_declareBuiltin___closed__5; LEAN_EXPORT lean_object* l_Lean_builtinInitAttr; +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_registerInitAttrUnsafe___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,7 +94,6 @@ lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208_(lean_object*); lean_object* lean_run_init(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_declareBuiltin___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___closed__3; static lean_object* l_Lean_declareBuiltin___closed__2; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__2; @@ -134,7 +135,6 @@ static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____cl lean_object* l_Lean_ParametricAttribute_setParam___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_runModInit___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__21; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_declareBuiltin___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_regularInitAttr; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__7; lean_object* l_Lean_Environment_addAndCompile(lean_object*, lean_object*, lean_object*); @@ -145,6 +145,7 @@ static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed_ extern lean_object* l_Lean_NameSet_empty; static lean_object* l___private_Lean_Compiler_InitAttr_0__Lean_isUnitType___closed__1; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_declareBuiltin___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -169,7 +170,6 @@ static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed_ extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___spec__7___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_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_InitAttr_0__Lean_isUnitType___boxed(lean_object*); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -189,8 +189,8 @@ static lean_object* l_Lean_registerInitAttrUnsafe___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_isIOUnitBuiltinInitFn___boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___lambda__1___closed__4; -LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_registerInitAttrUnsafe___spec__5(lean_object*, lean_object*); +lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_declareBuiltin___closed__6; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__22; lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); @@ -3461,45 +3461,7 @@ x_5 = l_Lean_ParametricAttribute_setParam___rarg(x_4, x_1, x_2, x_3); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_declareBuiltin___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 5); -x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_inc(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); -lean_ctor_set_tag(x_6, 1); -lean_ctor_set(x_6, 0, x_9); -return x_6; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_6, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_6); -lean_inc(x_5); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; -} -} -} -LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -3692,12 +3654,10 @@ lean_inc(x_21); lean_dec(x_20); x_22 = l_Lean_KernelException_toMessageData(x_21, x_8); x_23 = lean_ctor_get(x_3, 5); -lean_inc(x_23); x_24 = l_Lean_MessageData_toString(x_22, x_18); 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; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -lean_dec(x_23); x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_24, 1); @@ -3721,16 +3681,12 @@ x_34 = l_Lean_declareBuiltin___closed__13; 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_throwError___at_Lean_declareBuiltin___spec__1(x_35, x_3, x_4, x_26); -lean_dec(x_4); -lean_dec(x_3); +x_36 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_35, x_3, x_4, x_26); return x_36; } else { uint8_t x_37; -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); x_37 = !lean_is_exclusive(x_24); if (x_37 == 0) @@ -3742,6 +3698,7 @@ x_40 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_40, 0, x_39); x_41 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_41, 0, x_40); +lean_inc(x_23); x_42 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_42, 0, x_23); lean_ctor_set(x_42, 1, x_41); @@ -3761,6 +3718,7 @@ 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_inc(x_23); x_48 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_48, 0, x_23); lean_ctor_set(x_48, 1, x_47); @@ -3782,28 +3740,22 @@ x_51 = l_Lean_getBuiltinInitFnNameFor_x3f___closed__1; x_52 = lean_box(0); x_53 = l_Lean_ParametricAttribute_setParam___rarg(x_51, x_50, x_7, x_52); x_54 = lean_ctor_get(x_3, 5); -lean_inc(x_54); -x_55 = l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(x_53, x_18); +x_55 = l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(x_53, x_18); lean_dec(x_53); if (lean_obj_tag(x_55) == 0) { lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_54); x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); lean_dec(x_55); x_58 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_56, x_3, x_4, x_57); -lean_dec(x_4); -lean_dec(x_3); return x_58; } else { uint8_t x_59; -lean_dec(x_4); -lean_dec(x_3); x_59 = !lean_is_exclusive(x_55); if (x_59 == 0) { @@ -3814,6 +3766,7 @@ x_62 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_62, 0, x_61); x_63 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_63, 0, x_62); +lean_inc(x_54); x_64 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_64, 0, x_54); lean_ctor_set(x_64, 1, x_63); @@ -3833,6 +3786,7 @@ x_68 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_68, 0, x_67); x_69 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_69, 0, x_68); +lean_inc(x_54); x_70 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_70, 0, x_54); lean_ctor_set(x_70, 1, x_69); @@ -3845,23 +3799,23 @@ return x_71; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_declareBuiltin___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_declareBuiltin___spec__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_3; +x_3 = l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; } } -LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_declareBuiltin___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_3; -x_3 = l_IO_ofExcept___at_Lean_declareBuiltin___spec__2(x_1, x_2); -lean_dec(x_1); -return x_3; +lean_object* x_6; +x_6 = l_Lean_declareBuiltin(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_6; } } lean_object* initialize_Lean_Elab_InfoTree_Main(uint8_t builtin, lean_object*); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c index 1fcf29f71ea3..1439595f6ede 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/AuxDeclCache.c @@ -36,7 +36,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Compiler_LCN static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_cacheAuxDecl___spec__5___closed__1; lean_object* l_Lean_PersistentHashMap_instInhabited___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5341_(lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); 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; @@ -55,6 +54,7 @@ static lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___closed__3; +uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_cacheAuxDecl___closed__1; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, 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_5341_(x_5, x_9); +x_10 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341_(x_3, x_11); +x_13 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341_(x_3, x_17); +x_18 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341_(x_4, x_19); +x_21 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341_(x_4, x_26); +x_28 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341_(x_4, x_60); +x_63 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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 630a085e8dad..8b4130b26d56 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Basic.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Basic.c @@ -31,7 +31,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_visit___spec__2___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 uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5341____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg___boxed(lean_object*); uint8_t l_Lean_Compiler_LCNF_Code_isFun(lean_object*); @@ -42,6 +41,7 @@ lean_object* l_Lean_Compiler_LCNF_Decl_instantiateParamsLevelParams(lean_object* LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateJmpImp(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___lambda__2(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqImp(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_5389____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_instantiateParams(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedLetValue; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1(lean_object*); @@ -119,12 +119,12 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_markRecD size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetDecl____x40_Lean_Compiler_LCNF_Basic___hyg_1699____boxed(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_5389____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateTypeImp___closed__1; static lean_object* l_Lean_Compiler_LCNF_instInhabitedLetValue___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateUnreachImp___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectArgs___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_markRecDecls_go___spec__1(lean_object*, lean_object*, size_t, size_t, 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_5341____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_AltCore_forCodeM___boxed(lean_object*, lean_object*); uint8_t l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -148,6 +148,7 @@ static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_L static lean_object* l_Lean_Compiler_LCNF_instInhabitedCodeDecl___closed__1; LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_eqCases___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Param_toArg(lean_object*); +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectParams___spec__1(lean_object*, size_t, size_t, lean_object*); uint8_t l_Lean_Compiler_LCNF_Code_sizeLe(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1(lean_object*); @@ -169,12 +170,10 @@ uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1674_(l LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_instantiateRangeArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___spec__1___closed__2; lean_object* l_Lean_Compiler_LCNF_Decl_instantiateValueLevelParams_instCode(lean_object*, 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_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5341____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp___closed__1; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunImp___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5341_(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); lean_object* l_Lean_Compiler_LCNF_attachCodeDecls___boxed(lean_object*, lean_object*); @@ -262,7 +261,6 @@ lean_object* l_Lean_Compiler_LCNF_Code_isDecl___boxed(lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_LetValue_updateFVarImp___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_forM_go___spec__1___rarg___boxed(lean_object*, lean_object*, 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_5341____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__3; lean_object* l_Lean_Compiler_LCNF_Code_forM_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_833____boxed(lean_object*, lean_object*); @@ -287,6 +285,7 @@ uint64_t l_List_foldl___at_Lean_Expr_const___override___spec__1(uint64_t, lean_o lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedParam; lean_object* l_Lean_Compiler_LCNF_Code_sizeLe_inc___boxed(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_5389_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqLetValue____x40_Lean_Compiler_LCNF_Basic___hyg_833____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateLetImp___closed__1; static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___closed__1; @@ -297,6 +296,7 @@ lean_object* l_Lean_Compiler_LCNF_markRecDecls_visit___lambda__1___boxed(lean_ob lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_AltCore_getParams(lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_size(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_5389____spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Code_forM_go(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateFunDeclCoreImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instBEqLetDecl; @@ -322,6 +322,7 @@ LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Compiler_LCNF_Basi LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_instInhabitedLetDecl; lean_object* l_Lean_Compiler_LCNF_FunDeclCore_collectUsed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_collectType___spec__1___lambda__1___boxed(lean_object*); +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_5389____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqParam____x40_Lean_Compiler_LCNF_Basic___hyg_51____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_CasesCore_extractAlt_x21___closed__2; lean_object* l_Lean_Expr_instantiateLevelParamsNoCache(lean_object*, lean_object*, lean_object*); @@ -365,7 +366,6 @@ lean_object* l_Lean_Compiler_LCNF_Decl_instantiateParamsLevelParams___lambda__1( lean_object* l_Lean_Compiler_LCNF_Code_forM(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(lean_object*, lean_object*); static lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_updateAltsImp___closed__1; -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_5341____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Code_size_go___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_isTemplateLike___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_hasLocalInst___boxed(lean_object*); @@ -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_5341____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_5389____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_5341____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_5389____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_5341_(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_5389_(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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341____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_5389____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_5341_(x_1, x_2); +x_3 = l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_beqDecl____x40_Lean_Compiler_LCNF_Basic___hyg_5389_(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_5341____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_5389____boxed), 2, 0); return x_1; } } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Closure.c b/stage0/stdlib/Lean/Compiler/LCNF/Closure.c index e0dba73b4342..6b7ab5c140a9 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Closure.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Closure.c @@ -29,7 +29,7 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); lean_object* l_Lean_Compiler_LCNF_Closure_collectFVar(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_Compiler_LCNF_Closure_collectLetValue___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_Compiler_LCNF_Closure_collectType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___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_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1___closed__1; extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___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*); @@ -38,7 +38,7 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Compiler_LCNF size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1(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_Compiler_LCNF_Closure_collectParams___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Closure_collectLetValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -49,13 +49,14 @@ static lean_object* l_Lean_Compiler_LCNF_Closure_State_params___default___closed static lean_object* l_Lean_Compiler_LCNF_Closure_collectFVar___closed__4; lean_object* l_Lean_Compiler_LCNF_Closure_collectArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Closure_collectFVar___closed__1; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Closure_run___rarg___closed__1; static lean_object* l_Lean_Compiler_LCNF_Closure_collectType___closed__2; size_t lean_usize_mod(size_t, size_t); lean_object* l_Lean_Compiler_LCNF_Closure_run(lean_object*); extern lean_object* l_Lean_Compiler_LCNF_instMonadCompilerM; -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Closure_markVisited(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_Closure_collectFVar___closed__2; @@ -67,9 +68,11 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Closure_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_FVarIdSet_insert___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectLetValue___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_Compiler_LCNF_Closure_State_decls___default; size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Compiler_LCNF_Closure_collectType___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); @@ -219,7 +222,7 @@ uint8_t x_12; x_12 = lean_usize_dec_eq(x_2, x_3); 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_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_dec(x_4); x_13 = lean_array_uget(x_1, x_2); x_14 = lean_ctor_get(x_13, 2); @@ -227,70 +230,71 @@ lean_inc(x_14); lean_dec(x_13); x_15 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_16 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_17 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_17 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_15, x_16, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_15, x_16, x_14, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; size_t x_21; size_t x_22; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); -x_20 = 1; -x_21 = lean_usize_add(x_2, x_20); -x_2 = x_21; -x_4 = x_18; -x_11 = 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_2, x_21); +x_2 = x_22; +x_4 = x_19; +x_11 = x_20; goto _start; } 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); lean_dec(x_6); lean_dec(x_5); -x_23 = !lean_is_exclusive(x_17); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_18); +if (x_24 == 0) { -return x_17; +return x_18; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_17, 0); -x_25 = lean_ctor_get(x_17, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_18, 0); +x_26 = lean_ctor_get(x_18, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_17); -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_18); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } else { -lean_object* x_27; +lean_object* x_28; 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_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_4); -lean_ctor_set(x_27, 1, x_11); -return x_27; +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_4); +lean_ctor_set(x_28, 1, x_11); +return x_28; } } } @@ -726,409 +730,411 @@ return x_63; } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { -lean_dec(x_4); +lean_dec(x_5); switch (lean_obj_tag(x_1)) { case 5: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_1, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_1, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_1, 0); lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 1); +lean_inc(x_15); lean_dec(x_1); +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_3); lean_inc(x_2); -x_15 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); -return x_17; +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +return x_18; } else { -uint8_t x_18; -lean_dec(x_14); +uint8_t x_19; +lean_dec(x_15); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -return x_15; +return x_16; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -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; +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(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_dec(x_16); +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; } } } case 6: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_1, 1); -lean_inc(x_22); -x_23 = lean_ctor_get(x_1, 2); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_1, 1); lean_inc(x_23); +x_24 = lean_ctor_get(x_1, 2); +lean_inc(x_24); lean_dec(x_1); +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_3); lean_inc(x_2); -x_24 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_24) == 0) +x_25 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_23, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -lean_dec(x_24); -x_26 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_25); -return x_26; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_24, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_27; } else { -uint8_t x_27; -lean_dec(x_23); +uint8_t x_28; +lean_dec(x_24); +lean_dec(x_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_27 = !lean_is_exclusive(x_24); -if (x_27 == 0) +x_28 = !lean_is_exclusive(x_25); +if (x_28 == 0) { -return x_24; +return x_25; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_24, 0); -x_29 = lean_ctor_get(x_24, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_25, 0); +x_30 = lean_ctor_get(x_25, 1); +lean_inc(x_30); lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_24); -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_dec(x_25); +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; } } } case 7: { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_1, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_1, 2); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_1, 1); lean_inc(x_32); +x_33 = lean_ctor_get(x_1, 2); +lean_inc(x_33); lean_dec(x_1); +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_3); lean_inc(x_2); -x_33 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_33, 1); -lean_inc(x_34); -lean_dec(x_33); -x_35 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_34); -return x_35; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_33, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_35); +return x_36; } else { -uint8_t x_36; -lean_dec(x_32); +uint8_t x_37; +lean_dec(x_33); +lean_dec(x_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_33); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_34); +if (x_37 == 0) { -return x_33; +return x_34; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_33, 0); -x_38 = lean_ctor_get(x_33, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_34, 0); +x_39 = lean_ctor_get(x_34, 1); +lean_inc(x_39); lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_33); -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_dec(x_34); +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; } } } case 8: { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_40 = lean_ctor_get(x_1, 1); -lean_inc(x_40); -x_41 = lean_ctor_get(x_1, 2); +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_ctor_get(x_1, 1); lean_inc(x_41); -x_42 = lean_ctor_get(x_1, 3); +x_42 = lean_ctor_get(x_1, 2); lean_inc(x_42); +x_43 = lean_ctor_get(x_1, 3); +lean_inc(x_43); lean_dec(x_1); +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_3); lean_inc(x_2); -x_43 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_40, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_43) == 0) +x_44 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_41, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_43, 1); -lean_inc(x_44); -lean_dec(x_43); +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +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_3); lean_inc(x_2); -x_45 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_44); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_45, 1); -lean_inc(x_46); -lean_dec(x_45); -x_47 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_46); -return x_47; +x_46 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_43, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_47); +return x_48; } else { -uint8_t x_48; -lean_dec(x_42); +uint8_t x_49; +lean_dec(x_43); +lean_dec(x_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_48 = !lean_is_exclusive(x_45); -if (x_48 == 0) +x_49 = !lean_is_exclusive(x_46); +if (x_49 == 0) { -return x_45; +return x_46; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_45, 0); -x_50 = lean_ctor_get(x_45, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_46, 0); +x_51 = lean_ctor_get(x_46, 1); +lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_45); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_49); -lean_ctor_set(x_51, 1, x_50); -return x_51; +lean_dec(x_46); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_52; +uint8_t x_53; +lean_dec(x_43); lean_dec(x_42); -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); lean_dec(x_6); -lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_52 = !lean_is_exclusive(x_43); -if (x_52 == 0) +x_53 = !lean_is_exclusive(x_44); +if (x_53 == 0) { -return x_43; +return x_44; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_43, 0); -x_54 = lean_ctor_get(x_43, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_44, 0); +x_55 = lean_ctor_get(x_44, 1); +lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_43); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_dec(x_44); +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 10: { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_1, 1); -lean_inc(x_56); +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_1, 1); +lean_inc(x_57); lean_dec(x_1); -x_57 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_56, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_57; +x_58 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_57, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_58; } case 11: { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_1, 2); -lean_inc(x_58); +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_1, 2); +lean_inc(x_59); lean_dec(x_1); -x_59 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_58, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_59; +x_60 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_2, x_3, x_4, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_60; } default: { -lean_object* x_60; lean_object* x_61; +lean_object* x_61; lean_object* x_62; +lean_dec(x_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_60 = lean_box(0); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_12); -return x_61; +x_61 = lean_box(0); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_13); +return x_62; } } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___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_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__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) { _start: { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_inc(x_3); -x_12 = l_Lean_ForEachExprWhere_visited___at_Lean_Compiler_LCNF_Closure_collectType___spec__3(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_unbox(x_13); -lean_dec(x_13); -if (x_14 == 0) +lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_inc(x_4); +x_13 = l_Lean_ForEachExprWhere_visited___at_Lean_Compiler_LCNF_Closure_collectType___spec__3(x_4, 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_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_ctor_get(x_12, 1); -lean_inc(x_15); -lean_dec(x_12); +lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); lean_inc(x_1); -lean_inc(x_3); -x_16 = lean_apply_1(x_1, x_3); -x_17 = lean_unbox(x_16); -lean_dec(x_16); -if (x_17 == 0) +lean_inc(x_4); +x_17 = lean_apply_1(x_1, x_4); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_box(0); -x_19 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(x_3, x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -return x_19; +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(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_16); +return x_20; } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -lean_inc(x_3); -x_20 = l_Lean_ForEachExprWhere_checked___at_Lean_Compiler_LCNF_Closure_collectType___spec__4(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_15); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_unbox(x_21); -lean_dec(x_21); -if (x_22 == 0) +lean_object* x_21; lean_object* x_22; uint8_t x_23; +lean_inc(x_4); +x_21 = l_Lean_ForEachExprWhere_checked___at_Lean_Compiler_LCNF_Closure_collectType___spec__4(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_16); +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_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_21, 1); +lean_inc(x_24); +lean_dec(x_21); lean_inc(x_2); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_3); -x_24 = lean_apply_8(x_2, x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_23); -if (lean_obj_tag(x_24) == 0) +lean_inc(x_4); +x_25 = lean_apply_8(x_2, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_24); +if (lean_obj_tag(x_25) == 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); +if (x_3 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(x_3, x_1, x_2, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_26); -return x_27; +lean_dec(x_25); +x_27 = lean_box(0); +x_28 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(x_4, x_1, x_2, x_3, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +return x_28; } else { -uint8_t x_28; +uint8_t x_29; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1136,37 +1142,28 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_28 = !lean_is_exclusive(x_24); -if (x_28 == 0) -{ -return x_24; -} -else +x_29 = !lean_is_exclusive(x_25); +if (x_29 == 0) { -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_30; lean_object* x_31; +x_30 = lean_ctor_get(x_25, 0); +lean_dec(x_30); +x_31 = lean_box(0); +lean_ctor_set(x_25, 0, x_31); +return x_25; } else { lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_20, 1); +x_32 = lean_ctor_get(x_25, 1); lean_inc(x_32); -lean_dec(x_20); +lean_dec(x_25); x_33 = lean_box(0); -x_34 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(x_3, x_1, x_2, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_32); return x_34; } } @@ -1174,6 +1171,7 @@ return x_34; else { uint8_t x_35; +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1181,99 +1179,143 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_12); +x_35 = !lean_is_exclusive(x_25); if (x_35 == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_12, 0); -lean_dec(x_36); -x_37 = lean_box(0); -lean_ctor_set(x_12, 0, x_37); -return x_12; +return x_25; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_12, 1); -lean_inc(x_38); -lean_dec(x_12); -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; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_25, 0); +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_25); +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_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_21, 1); +lean_inc(x_39); +lean_dec(x_21); +x_40 = lean_box(0); +x_41 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(x_4, x_1, x_2, x_3, x_40, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_39); +return x_41; +} +} +} +else +{ +uint8_t x_42; +lean_dec(x_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_42 = !lean_is_exclusive(x_13); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_13, 0); +lean_dec(x_43); +x_44 = lean_box(0); +lean_ctor_set(x_13, 0, x_44); +return x_13; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_13, 1); +lean_inc(x_45); +lean_dec(x_13); +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; } } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___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_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_11 = l_Lean_ForEachExprWhere_initCache; -x_12 = lean_st_mk_ref(x_11, x_10); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = l_Lean_ForEachExprWhere_initCache; +x_13 = lean_st_mk_ref(x_12, x_11); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_dec(x_12); -lean_inc(x_13); -x_15 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_1, x_2, x_3, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_15) == 0) +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +lean_inc(x_14); +x_16 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2(x_1, x_2, x_4, x_3, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +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); -lean_dec(x_15); -x_18 = lean_st_ref_get(x_13, x_17); -lean_dec(x_13); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_st_ref_get(x_14, x_18); +lean_dec(x_14); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 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; +lean_object* x_21; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +lean_ctor_set(x_19, 0, x_17); +return x_19; } 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; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_17); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } else { -uint8_t x_23; -lean_dec(x_13); -x_23 = !lean_is_exclusive(x_15); -if (x_23 == 0) +uint8_t x_24; +lean_dec(x_14); +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; } } } @@ -1306,11 +1348,12 @@ return x_1; lean_object* l_Lean_Compiler_LCNF_Closure_collectType(lean_object* x_1, lean_object* x_2, 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_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; x_9 = l_Lean_Compiler_LCNF_Closure_collectType___closed__1; x_10 = l_Lean_Compiler_LCNF_Closure_collectType___closed__2; -x_11 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_9, x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; +x_11 = 0; +x_12 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_9, x_10, x_1, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_12; } } static lean_object* _init_l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1___closed__1() { @@ -1477,7 +1520,7 @@ return x_33; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; x_34 = lean_ctor_get(x_30, 0); lean_inc(x_34); lean_dec(x_30); @@ -1494,6 +1537,7 @@ x_39 = lean_ctor_get(x_34, 3); lean_inc(x_39); x_40 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_41 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_42 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); @@ -1501,141 +1545,141 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_38); -x_42 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_40, x_41, x_38, x_2, x_3, x_4, x_5, x_6, x_7, x_35); -if (lean_obj_tag(x_42) == 0) +x_43 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_40, x_41, x_38, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_35); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = lean_ctor_get(x_42, 1); -lean_inc(x_43); -lean_dec(x_42); -x_44 = lean_ctor_get(x_2, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_43, 1); lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_ctor_get(x_2, 1); +lean_inc(x_45); lean_inc(x_36); -x_45 = lean_apply_1(x_44, x_36); -x_46 = lean_unbox(x_45); -lean_dec(x_45); -if (x_46 == 0) +x_46 = lean_apply_1(x_45, x_36); +x_47 = lean_unbox(x_46); +lean_dec(x_46); +if (x_47 == 0) { -lean_object* x_47; +lean_object* x_48; lean_dec(x_38); lean_dec(x_37); lean_dec(x_36); lean_inc(x_3); -x_47 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_43); -if (lean_obj_tag(x_47) == 0) +x_48 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_44); +if (lean_obj_tag(x_48) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = lean_st_ref_take(x_3, x_48); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +x_49 = lean_ctor_get(x_48, 1); +lean_inc(x_49); +lean_dec(x_48); +x_50 = lean_st_ref_take(x_3, x_49); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_49); -x_52 = !lean_is_exclusive(x_50); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; -x_53 = lean_ctor_get(x_50, 2); -x_54 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_54, 0, x_34); -x_55 = lean_array_push(x_53, x_54); -lean_ctor_set(x_50, 2, x_55); -x_56 = lean_st_ref_set(x_3, x_50, x_51); +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = !lean_is_exclusive(x_51); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_54 = lean_ctor_get(x_51, 2); +x_55 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_55, 0, x_34); +x_56 = lean_array_push(x_54, x_55); +lean_ctor_set(x_51, 2, x_56); +x_57 = lean_st_ref_set(x_3, x_51, x_52); lean_dec(x_3); -x_57 = !lean_is_exclusive(x_56); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_56, 0); -lean_dec(x_58); -x_59 = lean_box(0); -lean_ctor_set(x_56, 0, x_59); -return x_56; +lean_object* x_59; lean_object* x_60; +x_59 = lean_ctor_get(x_57, 0); +lean_dec(x_59); +x_60 = lean_box(0); +lean_ctor_set(x_57, 0, x_60); +return x_57; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_56, 1); -lean_inc(x_60); -lean_dec(x_56); -x_61 = lean_box(0); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_60); -return x_62; +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +lean_dec(x_57); +x_62 = lean_box(0); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_61); +return x_63; } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; -x_63 = lean_ctor_get(x_50, 0); -x_64 = lean_ctor_get(x_50, 1); -x_65 = lean_ctor_get(x_50, 2); +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_64 = lean_ctor_get(x_51, 0); +x_65 = lean_ctor_get(x_51, 1); +x_66 = lean_ctor_get(x_51, 2); +lean_inc(x_66); lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_50); -x_66 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_66, 0, x_34); -x_67 = lean_array_push(x_65, x_66); -x_68 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_68, 0, x_63); -lean_ctor_set(x_68, 1, x_64); -lean_ctor_set(x_68, 2, x_67); -x_69 = lean_st_ref_set(x_3, x_68, x_51); +lean_dec(x_51); +x_67 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_67, 0, x_34); +x_68 = lean_array_push(x_66, x_67); +x_69 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_69, 0, x_64); +lean_ctor_set(x_69, 1, x_65); +lean_ctor_set(x_69, 2, x_68); +x_70 = lean_st_ref_set(x_3, x_69, x_52); lean_dec(x_3); -x_70 = lean_ctor_get(x_69, 1); -lean_inc(x_70); -if (lean_is_exclusive(x_69)) { - lean_ctor_release(x_69, 0); - lean_ctor_release(x_69, 1); - x_71 = x_69; +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_69); - x_71 = lean_box(0); + lean_dec_ref(x_70); + x_72 = lean_box(0); } -x_72 = lean_box(0); -if (lean_is_scalar(x_71)) { - x_73 = lean_alloc_ctor(0, 2, 0); +x_73 = lean_box(0); +if (lean_is_scalar(x_72)) { + x_74 = lean_alloc_ctor(0, 2, 0); } else { - x_73 = x_71; + x_74 = x_72; } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_70); -return x_73; +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_71); +return x_74; } } else { -uint8_t x_74; +uint8_t x_75; lean_dec(x_34); lean_dec(x_3); -x_74 = !lean_is_exclusive(x_47); -if (x_74 == 0) +x_75 = !lean_is_exclusive(x_48); +if (x_75 == 0) { -return x_47; +return x_48; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_75 = lean_ctor_get(x_47, 0); -x_76 = lean_ctor_get(x_47, 1); +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_48, 0); +x_77 = lean_ctor_get(x_48, 1); +lean_inc(x_77); lean_inc(x_76); -lean_inc(x_75); -lean_dec(x_47); -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_dec(x_48); +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; } } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; lean_dec(x_39); lean_dec(x_34); lean_dec(x_7); @@ -1643,26 +1687,25 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_78 = lean_st_ref_take(x_3, x_43); -x_79 = lean_ctor_get(x_78, 0); -lean_inc(x_79); -x_80 = lean_ctor_get(x_78, 1); +x_79 = lean_st_ref_take(x_3, x_44); +x_80 = lean_ctor_get(x_79, 0); lean_inc(x_80); -lean_dec(x_78); -x_81 = !lean_is_exclusive(x_79); -if (x_81 == 0) +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +lean_dec(x_79); +x_82 = !lean_is_exclusive(x_80); +if (x_82 == 0) { -lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_82 = lean_ctor_get(x_79, 1); -x_83 = 0; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_83 = lean_ctor_get(x_80, 1); x_84 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_84, 0, x_36); lean_ctor_set(x_84, 1, x_37); lean_ctor_set(x_84, 2, x_38); -lean_ctor_set_uint8(x_84, sizeof(void*)*3, x_83); -x_85 = lean_array_push(x_82, x_84); -lean_ctor_set(x_79, 1, x_85); -x_86 = lean_st_ref_set(x_3, x_79, x_80); +lean_ctor_set_uint8(x_84, sizeof(void*)*3, x_42); +x_85 = lean_array_push(x_83, x_84); +lean_ctor_set(x_80, 1, x_85); +x_86 = lean_st_ref_set(x_3, x_80, x_81); lean_dec(x_3); x_87 = !lean_is_exclusive(x_86); if (x_87 == 0) @@ -1689,52 +1732,51 @@ return x_92; } else { -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; -x_93 = lean_ctor_get(x_79, 0); -x_94 = lean_ctor_get(x_79, 1); -x_95 = lean_ctor_get(x_79, 2); +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_93 = lean_ctor_get(x_80, 0); +x_94 = lean_ctor_get(x_80, 1); +x_95 = lean_ctor_get(x_80, 2); lean_inc(x_95); lean_inc(x_94); lean_inc(x_93); -lean_dec(x_79); -x_96 = 0; -x_97 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_97, 0, x_36); -lean_ctor_set(x_97, 1, x_37); -lean_ctor_set(x_97, 2, x_38); -lean_ctor_set_uint8(x_97, sizeof(void*)*3, x_96); -x_98 = lean_array_push(x_94, x_97); -x_99 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_99, 0, x_93); -lean_ctor_set(x_99, 1, x_98); -lean_ctor_set(x_99, 2, x_95); -x_100 = lean_st_ref_set(x_3, x_99, x_80); +lean_dec(x_80); +x_96 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_96, 0, x_36); +lean_ctor_set(x_96, 1, x_37); +lean_ctor_set(x_96, 2, x_38); +lean_ctor_set_uint8(x_96, sizeof(void*)*3, x_42); +x_97 = lean_array_push(x_94, x_96); +x_98 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_98, 0, x_93); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set(x_98, 2, x_95); +x_99 = lean_st_ref_set(x_3, x_98, x_81); lean_dec(x_3); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -if (lean_is_exclusive(x_100)) { - lean_ctor_release(x_100, 0); - lean_ctor_release(x_100, 1); - x_102 = x_100; +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_100); - x_102 = lean_box(0); + lean_dec_ref(x_99); + x_101 = lean_box(0); } -x_103 = lean_box(0); -if (lean_is_scalar(x_102)) { - x_104 = lean_alloc_ctor(0, 2, 0); +x_102 = lean_box(0); +if (lean_is_scalar(x_101)) { + x_103 = lean_alloc_ctor(0, 2, 0); } else { - x_104 = x_102; + x_103 = x_101; } -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_101); -return x_104; +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_100); +return x_103; } } } else { -uint8_t x_105; +uint8_t x_104; lean_dec(x_39); lean_dec(x_38); lean_dec(x_37); @@ -1746,43 +1788,44 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_105 = !lean_is_exclusive(x_42); -if (x_105 == 0) +x_104 = !lean_is_exclusive(x_43); +if (x_104 == 0) { -return x_42; +return x_43; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_42, 0); -x_107 = lean_ctor_get(x_42, 1); -lean_inc(x_107); +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_43, 0); +x_106 = lean_ctor_get(x_43, 1); lean_inc(x_106); -lean_dec(x_42); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_inc(x_105); +lean_dec(x_43); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } } } } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; lean_object* x_114; lean_dec(x_1); -x_109 = lean_ctor_get(x_26, 1); -lean_inc(x_109); +x_108 = lean_ctor_get(x_26, 1); +lean_inc(x_108); lean_dec(x_26); -x_110 = lean_ctor_get(x_27, 0); -lean_inc(x_110); +x_109 = lean_ctor_get(x_27, 0); +lean_inc(x_109); lean_dec(x_27); -x_111 = lean_ctor_get(x_110, 2); -lean_inc(x_111); -x_112 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); -x_113 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_110 = lean_ctor_get(x_109, 2); +lean_inc(x_110); +x_111 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); +x_112 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_113 = 0; lean_inc(x_3); -x_114 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_112, x_113, x_111, x_2, x_3, x_4, x_5, x_6, x_7, x_109); +x_114 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_111, x_112, x_110, x_113, x_2, x_3, x_4, x_5, x_6, x_7, x_108); if (lean_obj_tag(x_114) == 0) { lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t x_119; @@ -1800,7 +1843,7 @@ if (x_119 == 0) { lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; x_120 = lean_ctor_get(x_117, 1); -x_121 = lean_array_push(x_120, x_110); +x_121 = lean_array_push(x_120, x_109); lean_ctor_set(x_117, 1, x_121); x_122 = lean_st_ref_set(x_3, x_117, x_118); lean_dec(x_3); @@ -1837,7 +1880,7 @@ lean_inc(x_131); lean_inc(x_130); lean_inc(x_129); lean_dec(x_117); -x_132 = lean_array_push(x_130, x_110); +x_132 = lean_array_push(x_130, x_109); x_133 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_133, 0, x_129); lean_ctor_set(x_133, 1, x_132); @@ -1868,7 +1911,7 @@ return x_138; else { uint8_t x_139; -lean_dec(x_110); +lean_dec(x_109); lean_dec(x_3); x_139 = !lean_is_exclusive(x_114); if (x_139 == 0) @@ -2199,7 +2242,7 @@ return x_225; } else { -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_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; uint8_t x_234; lean_object* x_235; x_226 = lean_ctor_get(x_222, 0); lean_inc(x_226); lean_dec(x_222); @@ -2216,6 +2259,7 @@ x_231 = lean_ctor_get(x_226, 3); lean_inc(x_231); x_232 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_233 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_234 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); @@ -2223,117 +2267,117 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_230); -x_234 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_232, x_233, x_230, x_2, x_3, x_4, x_5, x_6, x_7, x_227); -if (lean_obj_tag(x_234) == 0) -{ -lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; -x_235 = lean_ctor_get(x_234, 1); -lean_inc(x_235); -lean_dec(x_234); -x_236 = lean_ctor_get(x_2, 1); +x_235 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_232, x_233, x_230, x_234, x_2, x_3, x_4, x_5, x_6, x_7, x_227); +if (lean_obj_tag(x_235) == 0) +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; uint8_t x_239; +x_236 = lean_ctor_get(x_235, 1); lean_inc(x_236); +lean_dec(x_235); +x_237 = lean_ctor_get(x_2, 1); +lean_inc(x_237); lean_inc(x_228); -x_237 = lean_apply_1(x_236, x_228); -x_238 = lean_unbox(x_237); -lean_dec(x_237); -if (x_238 == 0) +x_238 = lean_apply_1(x_237, x_228); +x_239 = lean_unbox(x_238); +lean_dec(x_238); +if (x_239 == 0) { -lean_object* x_239; +lean_object* x_240; lean_dec(x_230); lean_dec(x_229); lean_dec(x_228); lean_inc(x_3); -x_239 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_231, x_2, x_3, x_4, x_5, x_6, x_7, x_235); -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; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; -x_240 = lean_ctor_get(x_239, 1); -lean_inc(x_240); -lean_dec(x_239); -x_241 = lean_st_ref_take(x_3, x_240); -x_242 = lean_ctor_get(x_241, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); +x_240 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_231, x_2, x_3, x_4, x_5, x_6, x_7, x_236); +if (lean_obj_tag(x_240) == 0) +{ +lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_241 = lean_ctor_get(x_240, 1); +lean_inc(x_241); +lean_dec(x_240); +x_242 = lean_st_ref_take(x_3, x_241); +x_243 = lean_ctor_get(x_242, 0); lean_inc(x_243); -lean_dec(x_241); -x_244 = lean_ctor_get(x_242, 0); +x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); -x_245 = lean_ctor_get(x_242, 1); +lean_dec(x_242); +x_245 = lean_ctor_get(x_243, 0); lean_inc(x_245); -x_246 = lean_ctor_get(x_242, 2); +x_246 = lean_ctor_get(x_243, 1); lean_inc(x_246); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - lean_ctor_release(x_242, 1); - lean_ctor_release(x_242, 2); - x_247 = x_242; +x_247 = lean_ctor_get(x_243, 2); +lean_inc(x_247); +if (lean_is_exclusive(x_243)) { + lean_ctor_release(x_243, 0); + lean_ctor_release(x_243, 1); + lean_ctor_release(x_243, 2); + x_248 = x_243; } else { - lean_dec_ref(x_242); - x_247 = lean_box(0); -} -x_248 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_248, 0, x_226); -x_249 = lean_array_push(x_246, x_248); -if (lean_is_scalar(x_247)) { - x_250 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_243); + x_248 = lean_box(0); +} +x_249 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_249, 0, x_226); +x_250 = lean_array_push(x_247, x_249); +if (lean_is_scalar(x_248)) { + x_251 = lean_alloc_ctor(0, 3, 0); } else { - x_250 = x_247; + x_251 = x_248; } -lean_ctor_set(x_250, 0, x_244); -lean_ctor_set(x_250, 1, x_245); -lean_ctor_set(x_250, 2, x_249); -x_251 = lean_st_ref_set(x_3, x_250, x_243); +lean_ctor_set(x_251, 0, x_245); +lean_ctor_set(x_251, 1, x_246); +lean_ctor_set(x_251, 2, x_250); +x_252 = lean_st_ref_set(x_3, x_251, x_244); lean_dec(x_3); -x_252 = lean_ctor_get(x_251, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_253 = x_251; +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_254 = x_252; } else { - lean_dec_ref(x_251); - x_253 = lean_box(0); + lean_dec_ref(x_252); + x_254 = lean_box(0); } -x_254 = lean_box(0); -if (lean_is_scalar(x_253)) { - x_255 = lean_alloc_ctor(0, 2, 0); +x_255 = lean_box(0); +if (lean_is_scalar(x_254)) { + x_256 = lean_alloc_ctor(0, 2, 0); } else { - x_255 = x_253; + x_256 = x_254; } -lean_ctor_set(x_255, 0, x_254); -lean_ctor_set(x_255, 1, x_252); -return x_255; +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_253); +return x_256; } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; +lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_dec(x_226); lean_dec(x_3); -x_256 = lean_ctor_get(x_239, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_239, 1); +x_257 = lean_ctor_get(x_240, 0); lean_inc(x_257); -if (lean_is_exclusive(x_239)) { - lean_ctor_release(x_239, 0); - lean_ctor_release(x_239, 1); - x_258 = x_239; +x_258 = lean_ctor_get(x_240, 1); +lean_inc(x_258); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_259 = x_240; } else { - lean_dec_ref(x_239); - x_258 = lean_box(0); + lean_dec_ref(x_240); + x_259 = lean_box(0); } -if (lean_is_scalar(x_258)) { - x_259 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_259)) { + x_260 = lean_alloc_ctor(1, 2, 0); } else { - x_259 = x_258; + x_260 = x_259; } -lean_ctor_set(x_259, 0, x_256); -lean_ctor_set(x_259, 1, x_257); -return x_259; +lean_ctor_set(x_260, 0, x_257); +lean_ctor_set(x_260, 1, x_258); +return x_260; } } else { -lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t 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_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_dec(x_231); lean_dec(x_226); lean_dec(x_7); @@ -2341,43 +2385,42 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_260 = lean_st_ref_take(x_3, x_235); -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -x_262 = lean_ctor_get(x_260, 1); +x_261 = lean_st_ref_take(x_3, x_236); +x_262 = lean_ctor_get(x_261, 0); lean_inc(x_262); -lean_dec(x_260); -x_263 = lean_ctor_get(x_261, 0); +x_263 = lean_ctor_get(x_261, 1); lean_inc(x_263); -x_264 = lean_ctor_get(x_261, 1); +lean_dec(x_261); +x_264 = lean_ctor_get(x_262, 0); lean_inc(x_264); -x_265 = lean_ctor_get(x_261, 2); +x_265 = lean_ctor_get(x_262, 1); lean_inc(x_265); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - lean_ctor_release(x_261, 2); - x_266 = x_261; +x_266 = lean_ctor_get(x_262, 2); +lean_inc(x_266); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + lean_ctor_release(x_262, 2); + x_267 = x_262; } else { - lean_dec_ref(x_261); - x_266 = lean_box(0); + lean_dec_ref(x_262); + x_267 = lean_box(0); } -x_267 = 0; x_268 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_268, 0, x_228); lean_ctor_set(x_268, 1, x_229); lean_ctor_set(x_268, 2, x_230); -lean_ctor_set_uint8(x_268, sizeof(void*)*3, x_267); -x_269 = lean_array_push(x_264, x_268); -if (lean_is_scalar(x_266)) { +lean_ctor_set_uint8(x_268, sizeof(void*)*3, x_234); +x_269 = lean_array_push(x_265, x_268); +if (lean_is_scalar(x_267)) { x_270 = lean_alloc_ctor(0, 3, 0); } else { - x_270 = x_266; + x_270 = x_267; } -lean_ctor_set(x_270, 0, x_263); +lean_ctor_set(x_270, 0, x_264); lean_ctor_set(x_270, 1, x_269); -lean_ctor_set(x_270, 2, x_265); -x_271 = lean_st_ref_set(x_3, x_270, x_262); +lean_ctor_set(x_270, 2, x_266); +x_271 = lean_st_ref_set(x_3, x_270, x_263); lean_dec(x_3); x_272 = lean_ctor_get(x_271, 1); lean_inc(x_272); @@ -2414,16 +2457,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_276 = lean_ctor_get(x_234, 0); +x_276 = lean_ctor_get(x_235, 0); lean_inc(x_276); -x_277 = lean_ctor_get(x_234, 1); +x_277 = lean_ctor_get(x_235, 1); lean_inc(x_277); -if (lean_is_exclusive(x_234)) { - lean_ctor_release(x_234, 0); - lean_ctor_release(x_234, 1); - x_278 = x_234; +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + lean_ctor_release(x_235, 1); + x_278 = x_235; } else { - lean_dec_ref(x_234); + lean_dec_ref(x_235); x_278 = lean_box(0); } if (lean_is_scalar(x_278)) { @@ -2439,7 +2482,7 @@ return x_279; } else { -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_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; uint8_t x_285; lean_object* x_286; lean_dec(x_1); x_280 = lean_ctor_get(x_218, 1); lean_inc(x_280); @@ -2451,278 +2494,279 @@ x_282 = lean_ctor_get(x_281, 2); lean_inc(x_282); x_283 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_284 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_285 = 0; lean_inc(x_3); -x_285 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_283, x_284, x_282, x_2, x_3, x_4, x_5, x_6, x_7, x_280); -if (lean_obj_tag(x_285) == 0) -{ -lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; -x_286 = lean_ctor_get(x_285, 1); -lean_inc(x_286); -lean_dec(x_285); -x_287 = lean_st_ref_take(x_3, x_286); -x_288 = lean_ctor_get(x_287, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_287, 1); +x_286 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_283, x_284, x_282, x_285, x_2, x_3, x_4, x_5, x_6, x_7, x_280); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; +x_287 = lean_ctor_get(x_286, 1); +lean_inc(x_287); +lean_dec(x_286); +x_288 = lean_st_ref_take(x_3, x_287); +x_289 = lean_ctor_get(x_288, 0); lean_inc(x_289); -lean_dec(x_287); -x_290 = lean_ctor_get(x_288, 0); +x_290 = lean_ctor_get(x_288, 1); lean_inc(x_290); -x_291 = lean_ctor_get(x_288, 1); +lean_dec(x_288); +x_291 = lean_ctor_get(x_289, 0); lean_inc(x_291); -x_292 = lean_ctor_get(x_288, 2); +x_292 = lean_ctor_get(x_289, 1); lean_inc(x_292); -if (lean_is_exclusive(x_288)) { - lean_ctor_release(x_288, 0); - lean_ctor_release(x_288, 1); - lean_ctor_release(x_288, 2); - x_293 = x_288; +x_293 = lean_ctor_get(x_289, 2); +lean_inc(x_293); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + lean_ctor_release(x_289, 2); + x_294 = x_289; } else { - lean_dec_ref(x_288); - x_293 = lean_box(0); + lean_dec_ref(x_289); + x_294 = lean_box(0); } -x_294 = lean_array_push(x_291, x_281); -if (lean_is_scalar(x_293)) { - x_295 = lean_alloc_ctor(0, 3, 0); +x_295 = lean_array_push(x_292, x_281); +if (lean_is_scalar(x_294)) { + x_296 = lean_alloc_ctor(0, 3, 0); } else { - x_295 = x_293; + x_296 = x_294; } -lean_ctor_set(x_295, 0, x_290); -lean_ctor_set(x_295, 1, x_294); -lean_ctor_set(x_295, 2, x_292); -x_296 = lean_st_ref_set(x_3, x_295, x_289); +lean_ctor_set(x_296, 0, x_291); +lean_ctor_set(x_296, 1, x_295); +lean_ctor_set(x_296, 2, x_293); +x_297 = lean_st_ref_set(x_3, x_296, x_290); lean_dec(x_3); -x_297 = lean_ctor_get(x_296, 1); -lean_inc(x_297); -if (lean_is_exclusive(x_296)) { - lean_ctor_release(x_296, 0); - lean_ctor_release(x_296, 1); - x_298 = x_296; +x_298 = lean_ctor_get(x_297, 1); +lean_inc(x_298); +if (lean_is_exclusive(x_297)) { + lean_ctor_release(x_297, 0); + lean_ctor_release(x_297, 1); + x_299 = x_297; } else { - lean_dec_ref(x_296); - x_298 = lean_box(0); + lean_dec_ref(x_297); + x_299 = lean_box(0); } -x_299 = lean_box(0); -if (lean_is_scalar(x_298)) { - x_300 = lean_alloc_ctor(0, 2, 0); +x_300 = lean_box(0); +if (lean_is_scalar(x_299)) { + x_301 = lean_alloc_ctor(0, 2, 0); } else { - x_300 = x_298; + x_301 = x_299; } -lean_ctor_set(x_300, 0, x_299); -lean_ctor_set(x_300, 1, x_297); -return x_300; +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_298); +return x_301; } else { -lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_dec(x_281); lean_dec(x_3); -x_301 = lean_ctor_get(x_285, 0); -lean_inc(x_301); -x_302 = lean_ctor_get(x_285, 1); +x_302 = lean_ctor_get(x_286, 0); lean_inc(x_302); -if (lean_is_exclusive(x_285)) { - lean_ctor_release(x_285, 0); - lean_ctor_release(x_285, 1); - x_303 = x_285; +x_303 = lean_ctor_get(x_286, 1); +lean_inc(x_303); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + lean_ctor_release(x_286, 1); + x_304 = x_286; } else { - lean_dec_ref(x_285); - x_303 = lean_box(0); + lean_dec_ref(x_286); + x_304 = lean_box(0); } -if (lean_is_scalar(x_303)) { - x_304 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_304)) { + x_305 = lean_alloc_ctor(1, 2, 0); } else { - x_304 = x_303; + x_305 = x_304; } -lean_ctor_set(x_304, 0, x_301); -lean_ctor_set(x_304, 1, x_302); -return x_304; +lean_ctor_set(x_305, 0, x_302); +lean_ctor_set(x_305, 1, x_303); +return x_305; } } } else { -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; uint8_t x_312; +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; uint8_t x_313; lean_dec(x_1); -x_305 = lean_ctor_get(x_216, 0); -lean_inc(x_305); -lean_dec(x_216); -x_306 = lean_ctor_get(x_215, 1); +x_306 = lean_ctor_get(x_216, 0); lean_inc(x_306); -lean_dec(x_215); -x_307 = lean_ctor_get(x_2, 1); +lean_dec(x_216); +x_307 = lean_ctor_get(x_215, 1); lean_inc(x_307); -x_308 = lean_ctor_get(x_305, 0); +lean_dec(x_215); +x_308 = lean_ctor_get(x_2, 1); lean_inc(x_308); -x_309 = lean_ctor_get(x_305, 1); +x_309 = lean_ctor_get(x_306, 0); lean_inc(x_309); -x_310 = lean_ctor_get(x_305, 3); +x_310 = lean_ctor_get(x_306, 1); lean_inc(x_310); -lean_inc(x_308); -x_311 = lean_apply_1(x_307, x_308); -x_312 = lean_unbox(x_311); -lean_dec(x_311); -if (x_312 == 0) +x_311 = lean_ctor_get(x_306, 3); +lean_inc(x_311); +lean_inc(x_309); +x_312 = lean_apply_1(x_308, x_309); +x_313 = lean_unbox(x_312); +lean_dec(x_312); +if (x_313 == 0) { -lean_object* x_313; +lean_object* x_314; +lean_dec(x_311); lean_dec(x_310); lean_dec(x_309); -lean_dec(x_308); lean_inc(x_3); -lean_inc(x_305); -x_313 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_305, x_2, x_3, x_4, x_5, x_6, x_7, x_306); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; -x_314 = lean_ctor_get(x_313, 1); -lean_inc(x_314); -lean_dec(x_313); -x_315 = lean_st_ref_take(x_3, x_314); -x_316 = lean_ctor_get(x_315, 0); -lean_inc(x_316); -x_317 = lean_ctor_get(x_315, 1); +lean_inc(x_306); +x_314 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_306, x_2, x_3, x_4, x_5, x_6, x_7, x_307); +if (lean_obj_tag(x_314) == 0) +{ +lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_315 = lean_ctor_get(x_314, 1); +lean_inc(x_315); +lean_dec(x_314); +x_316 = lean_st_ref_take(x_3, x_315); +x_317 = lean_ctor_get(x_316, 0); lean_inc(x_317); -lean_dec(x_315); -x_318 = lean_ctor_get(x_316, 0); +x_318 = lean_ctor_get(x_316, 1); lean_inc(x_318); -x_319 = lean_ctor_get(x_316, 1); +lean_dec(x_316); +x_319 = lean_ctor_get(x_317, 0); lean_inc(x_319); -x_320 = lean_ctor_get(x_316, 2); +x_320 = lean_ctor_get(x_317, 1); lean_inc(x_320); -if (lean_is_exclusive(x_316)) { - lean_ctor_release(x_316, 0); - lean_ctor_release(x_316, 1); - lean_ctor_release(x_316, 2); - x_321 = x_316; +x_321 = lean_ctor_get(x_317, 2); +lean_inc(x_321); +if (lean_is_exclusive(x_317)) { + lean_ctor_release(x_317, 0); + lean_ctor_release(x_317, 1); + lean_ctor_release(x_317, 2); + x_322 = x_317; } else { - lean_dec_ref(x_316); - x_321 = lean_box(0); -} -x_322 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_322, 0, x_305); -x_323 = lean_array_push(x_320, x_322); -if (lean_is_scalar(x_321)) { - x_324 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_317); + x_322 = lean_box(0); +} +x_323 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_323, 0, x_306); +x_324 = lean_array_push(x_321, x_323); +if (lean_is_scalar(x_322)) { + x_325 = lean_alloc_ctor(0, 3, 0); } else { - x_324 = x_321; + x_325 = x_322; } -lean_ctor_set(x_324, 0, x_318); -lean_ctor_set(x_324, 1, x_319); -lean_ctor_set(x_324, 2, x_323); -x_325 = lean_st_ref_set(x_3, x_324, x_317); +lean_ctor_set(x_325, 0, x_319); +lean_ctor_set(x_325, 1, x_320); +lean_ctor_set(x_325, 2, x_324); +x_326 = lean_st_ref_set(x_3, x_325, x_318); lean_dec(x_3); -x_326 = lean_ctor_get(x_325, 1); -lean_inc(x_326); -if (lean_is_exclusive(x_325)) { - lean_ctor_release(x_325, 0); - lean_ctor_release(x_325, 1); - x_327 = x_325; +x_327 = lean_ctor_get(x_326, 1); +lean_inc(x_327); +if (lean_is_exclusive(x_326)) { + lean_ctor_release(x_326, 0); + lean_ctor_release(x_326, 1); + x_328 = x_326; } else { - lean_dec_ref(x_325); - x_327 = lean_box(0); + lean_dec_ref(x_326); + x_328 = lean_box(0); } -x_328 = lean_box(0); -if (lean_is_scalar(x_327)) { - x_329 = lean_alloc_ctor(0, 2, 0); +x_329 = lean_box(0); +if (lean_is_scalar(x_328)) { + x_330 = lean_alloc_ctor(0, 2, 0); } else { - x_329 = x_327; + x_330 = x_328; } -lean_ctor_set(x_329, 0, x_328); -lean_ctor_set(x_329, 1, x_326); -return x_329; +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_327); +return x_330; } else { -lean_object* x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; -lean_dec(x_305); +lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; +lean_dec(x_306); lean_dec(x_3); -x_330 = lean_ctor_get(x_313, 0); -lean_inc(x_330); -x_331 = lean_ctor_get(x_313, 1); +x_331 = lean_ctor_get(x_314, 0); lean_inc(x_331); -if (lean_is_exclusive(x_313)) { - lean_ctor_release(x_313, 0); - lean_ctor_release(x_313, 1); - x_332 = x_313; +x_332 = lean_ctor_get(x_314, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_314)) { + lean_ctor_release(x_314, 0); + lean_ctor_release(x_314, 1); + x_333 = x_314; } else { - lean_dec_ref(x_313); - x_332 = lean_box(0); + lean_dec_ref(x_314); + x_333 = lean_box(0); } -if (lean_is_scalar(x_332)) { - x_333 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_333)) { + x_334 = lean_alloc_ctor(1, 2, 0); } else { - x_333 = x_332; + x_334 = x_333; } -lean_ctor_set(x_333, 0, x_330); -lean_ctor_set(x_333, 1, x_331); -return x_333; +lean_ctor_set(x_334, 0, x_331); +lean_ctor_set(x_334, 1, x_332); +return x_334; } } 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; uint8_t 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_dec(x_305); +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; uint8_t x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; +lean_dec(x_306); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_334 = lean_st_ref_take(x_3, x_306); -x_335 = lean_ctor_get(x_334, 0); -lean_inc(x_335); -x_336 = lean_ctor_get(x_334, 1); +x_335 = lean_st_ref_take(x_3, x_307); +x_336 = lean_ctor_get(x_335, 0); lean_inc(x_336); -lean_dec(x_334); -x_337 = lean_ctor_get(x_335, 0); +x_337 = lean_ctor_get(x_335, 1); lean_inc(x_337); -x_338 = lean_ctor_get(x_335, 1); +lean_dec(x_335); +x_338 = lean_ctor_get(x_336, 0); lean_inc(x_338); -x_339 = lean_ctor_get(x_335, 2); +x_339 = lean_ctor_get(x_336, 1); lean_inc(x_339); -if (lean_is_exclusive(x_335)) { - lean_ctor_release(x_335, 0); - lean_ctor_release(x_335, 1); - lean_ctor_release(x_335, 2); - x_340 = x_335; +x_340 = lean_ctor_get(x_336, 2); +lean_inc(x_340); +if (lean_is_exclusive(x_336)) { + lean_ctor_release(x_336, 0); + lean_ctor_release(x_336, 1); + lean_ctor_release(x_336, 2); + x_341 = x_336; } else { - lean_dec_ref(x_335); - x_340 = lean_box(0); -} -x_341 = 0; -x_342 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_342, 0, x_308); -lean_ctor_set(x_342, 1, x_309); -lean_ctor_set(x_342, 2, x_310); -lean_ctor_set_uint8(x_342, sizeof(void*)*3, x_341); -x_343 = lean_array_push(x_338, x_342); -if (lean_is_scalar(x_340)) { - x_344 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_336); + x_341 = lean_box(0); +} +x_342 = 0; +x_343 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_343, 0, x_309); +lean_ctor_set(x_343, 1, x_310); +lean_ctor_set(x_343, 2, x_311); +lean_ctor_set_uint8(x_343, sizeof(void*)*3, x_342); +x_344 = lean_array_push(x_339, x_343); +if (lean_is_scalar(x_341)) { + x_345 = lean_alloc_ctor(0, 3, 0); } else { - x_344 = x_340; + x_345 = x_341; } -lean_ctor_set(x_344, 0, x_337); -lean_ctor_set(x_344, 1, x_343); -lean_ctor_set(x_344, 2, x_339); -x_345 = lean_st_ref_set(x_3, x_344, x_336); +lean_ctor_set(x_345, 0, x_338); +lean_ctor_set(x_345, 1, x_344); +lean_ctor_set(x_345, 2, x_340); +x_346 = lean_st_ref_set(x_3, x_345, x_337); lean_dec(x_3); -x_346 = lean_ctor_get(x_345, 1); -lean_inc(x_346); -if (lean_is_exclusive(x_345)) { - lean_ctor_release(x_345, 0); - lean_ctor_release(x_345, 1); - x_347 = x_345; +x_347 = lean_ctor_get(x_346, 1); +lean_inc(x_347); +if (lean_is_exclusive(x_346)) { + lean_ctor_release(x_346, 0); + lean_ctor_release(x_346, 1); + x_348 = x_346; } else { - lean_dec_ref(x_345); - x_347 = lean_box(0); + lean_dec_ref(x_346); + x_348 = lean_box(0); } -x_348 = lean_box(0); -if (lean_is_scalar(x_347)) { - x_349 = lean_alloc_ctor(0, 2, 0); +x_349 = lean_box(0); +if (lean_is_scalar(x_348)) { + x_350 = lean_alloc_ctor(0, 2, 0); } else { - x_349 = x_347; + x_350 = x_348; } -lean_ctor_set(x_349, 0, x_348); -lean_ctor_set(x_349, 1, x_346); -return x_349; +lean_ctor_set(x_350, 0, x_349); +lean_ctor_set(x_350, 1, x_347); +return x_350; } } } @@ -2730,7 +2774,7 @@ return x_349; } else { -lean_object* x_350; +lean_object* x_351; lean_dec(x_14); lean_dec(x_7); lean_dec(x_6); @@ -2739,48 +2783,48 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_350 = lean_box(0); -lean_ctor_set(x_9, 0, x_350); +x_351 = lean_box(0); +lean_ctor_set(x_9, 0, x_351); return x_9; } } else { -lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; -x_351 = lean_ctor_get(x_9, 0); -x_352 = lean_ctor_get(x_9, 1); +lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +x_352 = lean_ctor_get(x_9, 0); +x_353 = lean_ctor_get(x_9, 1); +lean_inc(x_353); lean_inc(x_352); -lean_inc(x_351); lean_dec(x_9); -x_353 = lean_ctor_get(x_351, 0); -lean_inc(x_353); -lean_dec(x_351); -x_354 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_353, x_1); -lean_dec(x_353); -if (lean_obj_tag(x_354) == 0) +x_354 = lean_ctor_get(x_352, 0); +lean_inc(x_354); +lean_dec(x_352); +x_355 = l_Lean_RBNode_findCore___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__2(x_354, x_1); +lean_dec(x_354); +if (lean_obj_tag(x_355) == 0) { -lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; uint8_t x_360; +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; uint8_t x_361; lean_inc(x_1); -x_355 = l_Lean_Compiler_LCNF_Closure_markVisited(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_352); -x_356 = lean_ctor_get(x_355, 1); -lean_inc(x_356); -if (lean_is_exclusive(x_355)) { - lean_ctor_release(x_355, 0); - lean_ctor_release(x_355, 1); - x_357 = x_355; +x_356 = l_Lean_Compiler_LCNF_Closure_markVisited(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_353); +x_357 = lean_ctor_get(x_356, 1); +lean_inc(x_357); +if (lean_is_exclusive(x_356)) { + lean_ctor_release(x_356, 0); + lean_ctor_release(x_356, 1); + x_358 = x_356; } else { - lean_dec_ref(x_355); - x_357 = lean_box(0); + lean_dec_ref(x_356); + x_358 = lean_box(0); } -x_358 = lean_ctor_get(x_2, 0); -lean_inc(x_358); +x_359 = lean_ctor_get(x_2, 0); +lean_inc(x_359); lean_inc(x_1); -x_359 = lean_apply_1(x_358, x_1); -x_360 = lean_unbox(x_359); -lean_dec(x_359); -if (x_360 == 0) +x_360 = lean_apply_1(x_359, x_1); +x_361 = lean_unbox(x_360); +lean_dec(x_360); +if (x_361 == 0) { -lean_object* x_361; lean_object* x_362; +lean_object* x_362; lean_object* x_363; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -2788,587 +2832,588 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_361 = lean_box(0); -if (lean_is_scalar(x_357)) { - x_362 = lean_alloc_ctor(0, 2, 0); +x_362 = lean_box(0); +if (lean_is_scalar(x_358)) { + x_363 = lean_alloc_ctor(0, 2, 0); } else { - x_362 = x_357; + x_363 = x_358; } -lean_ctor_set(x_362, 0, x_361); -lean_ctor_set(x_362, 1, x_356); -return x_362; +lean_ctor_set(x_363, 0, x_362); +lean_ctor_set(x_363, 1, x_357); +return x_363; } else { -lean_object* x_363; lean_object* x_364; -lean_dec(x_357); +lean_object* x_364; lean_object* x_365; +lean_dec(x_358); lean_inc(x_1); -x_363 = l_Lean_Compiler_LCNF_findFunDecl_x3f(x_1, x_4, x_5, x_6, x_7, x_356); -x_364 = lean_ctor_get(x_363, 0); -lean_inc(x_364); -if (lean_obj_tag(x_364) == 0) -{ -lean_object* x_365; lean_object* x_366; lean_object* x_367; -x_365 = lean_ctor_get(x_363, 1); +x_364 = l_Lean_Compiler_LCNF_findFunDecl_x3f(x_1, x_4, x_5, x_6, x_7, x_357); +x_365 = lean_ctor_get(x_364, 0); lean_inc(x_365); -lean_dec(x_363); -lean_inc(x_1); -x_366 = l_Lean_Compiler_LCNF_findParam_x3f(x_1, x_4, x_5, x_6, x_7, x_365); -x_367 = lean_ctor_get(x_366, 0); -lean_inc(x_367); -if (lean_obj_tag(x_367) == 0) +if (lean_obj_tag(x_365) == 0) { -lean_object* x_368; lean_object* x_369; lean_object* x_370; -x_368 = lean_ctor_get(x_366, 1); +lean_object* x_366; lean_object* x_367; lean_object* x_368; +x_366 = lean_ctor_get(x_364, 1); +lean_inc(x_366); +lean_dec(x_364); +lean_inc(x_1); +x_367 = l_Lean_Compiler_LCNF_findParam_x3f(x_1, x_4, x_5, x_6, x_7, x_366); +x_368 = lean_ctor_get(x_367, 0); lean_inc(x_368); -lean_dec(x_366); -x_369 = l_Lean_Compiler_LCNF_findLetDecl_x3f(x_1, x_4, x_5, x_6, x_7, x_368); -x_370 = lean_ctor_get(x_369, 0); -lean_inc(x_370); -if (lean_obj_tag(x_370) == 0) -{ -lean_object* x_371; lean_object* x_372; lean_object* x_373; -x_371 = lean_ctor_get(x_369, 1); +if (lean_obj_tag(x_368) == 0) +{ +lean_object* x_369; lean_object* x_370; lean_object* x_371; +x_369 = lean_ctor_get(x_367, 1); +lean_inc(x_369); +lean_dec(x_367); +x_370 = l_Lean_Compiler_LCNF_findLetDecl_x3f(x_1, x_4, x_5, x_6, x_7, x_369); +x_371 = lean_ctor_get(x_370, 0); lean_inc(x_371); -lean_dec(x_369); -x_372 = l_Lean_Compiler_LCNF_Closure_collectFVar___closed__4; -x_373 = l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1(x_372, x_2, x_3, x_4, x_5, x_6, x_7, x_371); -return x_373; +if (lean_obj_tag(x_371) == 0) +{ +lean_object* x_372; lean_object* x_373; lean_object* x_374; +x_372 = lean_ctor_get(x_370, 1); +lean_inc(x_372); +lean_dec(x_370); +x_373 = l_Lean_Compiler_LCNF_Closure_collectFVar___closed__4; +x_374 = l_panic___at_Lean_Compiler_LCNF_Closure_collectFVar___spec__1(x_373, x_2, x_3, x_4, x_5, x_6, x_7, x_372); +return x_374; } 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; lean_object* x_381; lean_object* x_382; -x_374 = lean_ctor_get(x_370, 0); -lean_inc(x_374); -lean_dec(x_370); -x_375 = lean_ctor_get(x_369, 1); +lean_object* x_375; lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; uint8_t x_383; lean_object* x_384; +x_375 = lean_ctor_get(x_371, 0); lean_inc(x_375); -lean_dec(x_369); -x_376 = lean_ctor_get(x_374, 0); +lean_dec(x_371); +x_376 = lean_ctor_get(x_370, 1); lean_inc(x_376); -x_377 = lean_ctor_get(x_374, 1); +lean_dec(x_370); +x_377 = lean_ctor_get(x_375, 0); lean_inc(x_377); -x_378 = lean_ctor_get(x_374, 2); +x_378 = lean_ctor_get(x_375, 1); lean_inc(x_378); -x_379 = lean_ctor_get(x_374, 3); +x_379 = lean_ctor_get(x_375, 2); lean_inc(x_379); -x_380 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); -x_381 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_380 = lean_ctor_get(x_375, 3); +lean_inc(x_380); +x_381 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); +x_382 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_383 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_378); -x_382 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_380, x_381, x_378, x_2, x_3, x_4, x_5, x_6, x_7, x_375); -if (lean_obj_tag(x_382) == 0) -{ -lean_object* x_383; lean_object* x_384; lean_object* x_385; uint8_t x_386; -x_383 = lean_ctor_get(x_382, 1); -lean_inc(x_383); -lean_dec(x_382); -x_384 = lean_ctor_get(x_2, 1); -lean_inc(x_384); -lean_inc(x_376); -x_385 = lean_apply_1(x_384, x_376); -x_386 = lean_unbox(x_385); -lean_dec(x_385); -if (x_386 == 0) +lean_inc(x_379); +x_384 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_381, x_382, x_379, x_383, x_2, x_3, x_4, x_5, x_6, x_7, x_376); +if (lean_obj_tag(x_384) == 0) +{ +lean_object* x_385; lean_object* x_386; lean_object* x_387; uint8_t x_388; +x_385 = lean_ctor_get(x_384, 1); +lean_inc(x_385); +lean_dec(x_384); +x_386 = lean_ctor_get(x_2, 1); +lean_inc(x_386); +lean_inc(x_377); +x_387 = lean_apply_1(x_386, x_377); +x_388 = lean_unbox(x_387); +lean_dec(x_387); +if (x_388 == 0) { -lean_object* x_387; +lean_object* x_389; +lean_dec(x_379); lean_dec(x_378); lean_dec(x_377); -lean_dec(x_376); lean_inc(x_3); -x_387 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_379, x_2, x_3, x_4, x_5, x_6, x_7, x_383); -if (lean_obj_tag(x_387) == 0) +x_389 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_380, x_2, x_3, x_4, x_5, x_6, x_7, x_385); +if (lean_obj_tag(x_389) == 0) { -lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; -x_388 = lean_ctor_get(x_387, 1); -lean_inc(x_388); -lean_dec(x_387); -x_389 = lean_st_ref_take(x_3, x_388); -x_390 = lean_ctor_get(x_389, 0); +lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; 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_390 = lean_ctor_get(x_389, 1); lean_inc(x_390); -x_391 = lean_ctor_get(x_389, 1); -lean_inc(x_391); lean_dec(x_389); -x_392 = lean_ctor_get(x_390, 0); +x_391 = lean_st_ref_take(x_3, x_390); +x_392 = lean_ctor_get(x_391, 0); lean_inc(x_392); -x_393 = lean_ctor_get(x_390, 1); +x_393 = lean_ctor_get(x_391, 1); lean_inc(x_393); -x_394 = lean_ctor_get(x_390, 2); +lean_dec(x_391); +x_394 = lean_ctor_get(x_392, 0); lean_inc(x_394); -if (lean_is_exclusive(x_390)) { - lean_ctor_release(x_390, 0); - lean_ctor_release(x_390, 1); - lean_ctor_release(x_390, 2); - x_395 = x_390; +x_395 = lean_ctor_get(x_392, 1); +lean_inc(x_395); +x_396 = lean_ctor_get(x_392, 2); +lean_inc(x_396); +if (lean_is_exclusive(x_392)) { + lean_ctor_release(x_392, 0); + lean_ctor_release(x_392, 1); + lean_ctor_release(x_392, 2); + x_397 = x_392; } else { - lean_dec_ref(x_390); - x_395 = lean_box(0); -} -x_396 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_396, 0, x_374); -x_397 = lean_array_push(x_394, x_396); -if (lean_is_scalar(x_395)) { - x_398 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_392); + x_397 = lean_box(0); +} +x_398 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_398, 0, x_375); +x_399 = lean_array_push(x_396, x_398); +if (lean_is_scalar(x_397)) { + x_400 = lean_alloc_ctor(0, 3, 0); } else { - x_398 = x_395; + x_400 = x_397; } -lean_ctor_set(x_398, 0, x_392); -lean_ctor_set(x_398, 1, x_393); -lean_ctor_set(x_398, 2, x_397); -x_399 = lean_st_ref_set(x_3, x_398, x_391); +lean_ctor_set(x_400, 0, x_394); +lean_ctor_set(x_400, 1, x_395); +lean_ctor_set(x_400, 2, x_399); +x_401 = lean_st_ref_set(x_3, x_400, x_393); lean_dec(x_3); -x_400 = lean_ctor_get(x_399, 1); -lean_inc(x_400); -if (lean_is_exclusive(x_399)) { - lean_ctor_release(x_399, 0); - lean_ctor_release(x_399, 1); - x_401 = x_399; +x_402 = lean_ctor_get(x_401, 1); +lean_inc(x_402); +if (lean_is_exclusive(x_401)) { + lean_ctor_release(x_401, 0); + lean_ctor_release(x_401, 1); + x_403 = x_401; } else { - lean_dec_ref(x_399); - x_401 = lean_box(0); + lean_dec_ref(x_401); + x_403 = lean_box(0); } -x_402 = lean_box(0); -if (lean_is_scalar(x_401)) { - x_403 = lean_alloc_ctor(0, 2, 0); +x_404 = lean_box(0); +if (lean_is_scalar(x_403)) { + x_405 = lean_alloc_ctor(0, 2, 0); } else { - x_403 = x_401; + x_405 = x_403; } -lean_ctor_set(x_403, 0, x_402); -lean_ctor_set(x_403, 1, x_400); -return x_403; +lean_ctor_set(x_405, 0, x_404); +lean_ctor_set(x_405, 1, x_402); +return x_405; } else { -lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; -lean_dec(x_374); +lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; +lean_dec(x_375); lean_dec(x_3); -x_404 = lean_ctor_get(x_387, 0); -lean_inc(x_404); -x_405 = lean_ctor_get(x_387, 1); -lean_inc(x_405); -if (lean_is_exclusive(x_387)) { - lean_ctor_release(x_387, 0); - lean_ctor_release(x_387, 1); - x_406 = x_387; +x_406 = lean_ctor_get(x_389, 0); +lean_inc(x_406); +x_407 = lean_ctor_get(x_389, 1); +lean_inc(x_407); +if (lean_is_exclusive(x_389)) { + lean_ctor_release(x_389, 0); + lean_ctor_release(x_389, 1); + x_408 = x_389; } else { - lean_dec_ref(x_387); - x_406 = lean_box(0); + lean_dec_ref(x_389); + x_408 = lean_box(0); } -if (lean_is_scalar(x_406)) { - x_407 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_408)) { + x_409 = lean_alloc_ctor(1, 2, 0); } else { - x_407 = x_406; + x_409 = x_408; } -lean_ctor_set(x_407, 0, x_404); -lean_ctor_set(x_407, 1, x_405); -return x_407; +lean_ctor_set(x_409, 0, x_406); +lean_ctor_set(x_409, 1, x_407); +return x_409; } } else { -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; uint8_t x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; -lean_dec(x_379); -lean_dec(x_374); +lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; +lean_dec(x_380); +lean_dec(x_375); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_408 = lean_st_ref_take(x_3, x_383); -x_409 = lean_ctor_get(x_408, 0); -lean_inc(x_409); -x_410 = lean_ctor_get(x_408, 1); -lean_inc(x_410); -lean_dec(x_408); -x_411 = lean_ctor_get(x_409, 0); +x_410 = lean_st_ref_take(x_3, x_385); +x_411 = lean_ctor_get(x_410, 0); lean_inc(x_411); -x_412 = lean_ctor_get(x_409, 1); +x_412 = lean_ctor_get(x_410, 1); lean_inc(x_412); -x_413 = lean_ctor_get(x_409, 2); +lean_dec(x_410); +x_413 = lean_ctor_get(x_411, 0); lean_inc(x_413); -if (lean_is_exclusive(x_409)) { - lean_ctor_release(x_409, 0); - lean_ctor_release(x_409, 1); - lean_ctor_release(x_409, 2); - x_414 = x_409; +x_414 = lean_ctor_get(x_411, 1); +lean_inc(x_414); +x_415 = lean_ctor_get(x_411, 2); +lean_inc(x_415); +if (lean_is_exclusive(x_411)) { + lean_ctor_release(x_411, 0); + lean_ctor_release(x_411, 1); + lean_ctor_release(x_411, 2); + x_416 = x_411; } else { - lean_dec_ref(x_409); - x_414 = lean_box(0); -} -x_415 = 0; -x_416 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_416, 0, x_376); -lean_ctor_set(x_416, 1, x_377); -lean_ctor_set(x_416, 2, x_378); -lean_ctor_set_uint8(x_416, sizeof(void*)*3, x_415); -x_417 = lean_array_push(x_412, x_416); -if (lean_is_scalar(x_414)) { - x_418 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_411); + x_416 = lean_box(0); +} +x_417 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_417, 0, x_377); +lean_ctor_set(x_417, 1, x_378); +lean_ctor_set(x_417, 2, x_379); +lean_ctor_set_uint8(x_417, sizeof(void*)*3, x_383); +x_418 = lean_array_push(x_414, x_417); +if (lean_is_scalar(x_416)) { + x_419 = lean_alloc_ctor(0, 3, 0); } else { - x_418 = x_414; + x_419 = x_416; } -lean_ctor_set(x_418, 0, x_411); -lean_ctor_set(x_418, 1, x_417); -lean_ctor_set(x_418, 2, x_413); -x_419 = lean_st_ref_set(x_3, x_418, x_410); +lean_ctor_set(x_419, 0, x_413); +lean_ctor_set(x_419, 1, x_418); +lean_ctor_set(x_419, 2, x_415); +x_420 = lean_st_ref_set(x_3, x_419, x_412); lean_dec(x_3); -x_420 = lean_ctor_get(x_419, 1); -lean_inc(x_420); -if (lean_is_exclusive(x_419)) { - lean_ctor_release(x_419, 0); - lean_ctor_release(x_419, 1); - x_421 = x_419; +x_421 = lean_ctor_get(x_420, 1); +lean_inc(x_421); +if (lean_is_exclusive(x_420)) { + lean_ctor_release(x_420, 0); + lean_ctor_release(x_420, 1); + x_422 = x_420; } else { - lean_dec_ref(x_419); - x_421 = lean_box(0); + lean_dec_ref(x_420); + x_422 = lean_box(0); } -x_422 = lean_box(0); -if (lean_is_scalar(x_421)) { - x_423 = lean_alloc_ctor(0, 2, 0); +x_423 = lean_box(0); +if (lean_is_scalar(x_422)) { + x_424 = lean_alloc_ctor(0, 2, 0); } else { - x_423 = x_421; + x_424 = x_422; } -lean_ctor_set(x_423, 0, x_422); -lean_ctor_set(x_423, 1, x_420); -return x_423; +lean_ctor_set(x_424, 0, x_423); +lean_ctor_set(x_424, 1, x_421); +return x_424; } } else { -lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; +lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; +lean_dec(x_380); lean_dec(x_379); lean_dec(x_378); lean_dec(x_377); -lean_dec(x_376); -lean_dec(x_374); +lean_dec(x_375); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_424 = lean_ctor_get(x_382, 0); -lean_inc(x_424); -x_425 = lean_ctor_get(x_382, 1); +x_425 = lean_ctor_get(x_384, 0); lean_inc(x_425); -if (lean_is_exclusive(x_382)) { - lean_ctor_release(x_382, 0); - lean_ctor_release(x_382, 1); - x_426 = x_382; +x_426 = lean_ctor_get(x_384, 1); +lean_inc(x_426); +if (lean_is_exclusive(x_384)) { + lean_ctor_release(x_384, 0); + lean_ctor_release(x_384, 1); + x_427 = x_384; } else { - lean_dec_ref(x_382); - x_426 = lean_box(0); + lean_dec_ref(x_384); + x_427 = lean_box(0); } -if (lean_is_scalar(x_426)) { - x_427 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_427)) { + x_428 = lean_alloc_ctor(1, 2, 0); } else { - x_427 = x_426; + x_428 = x_427; } -lean_ctor_set(x_427, 0, x_424); -lean_ctor_set(x_427, 1, x_425); -return x_427; +lean_ctor_set(x_428, 0, x_425); +lean_ctor_set(x_428, 1, x_426); +return x_428; } } } else { -lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; +lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; uint8_t x_434; lean_object* x_435; lean_dec(x_1); -x_428 = lean_ctor_get(x_366, 1); -lean_inc(x_428); -lean_dec(x_366); -x_429 = lean_ctor_get(x_367, 0); +x_429 = lean_ctor_get(x_367, 1); lean_inc(x_429); lean_dec(x_367); -x_430 = lean_ctor_get(x_429, 2); +x_430 = lean_ctor_get(x_368, 0); lean_inc(x_430); -x_431 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); -x_432 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +lean_dec(x_368); +x_431 = lean_ctor_get(x_430, 2); +lean_inc(x_431); +x_432 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); +x_433 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_434 = 0; lean_inc(x_3); -x_433 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_431, x_432, x_430, x_2, x_3, x_4, x_5, x_6, x_7, x_428); -if (lean_obj_tag(x_433) == 0) -{ -lean_object* x_434; lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; -x_434 = lean_ctor_get(x_433, 1); -lean_inc(x_434); -lean_dec(x_433); -x_435 = lean_st_ref_take(x_3, x_434); -x_436 = lean_ctor_get(x_435, 0); +x_435 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_432, x_433, x_431, x_434, x_2, x_3, x_4, x_5, x_6, x_7, x_429); +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; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; lean_object* x_448; lean_object* x_449; lean_object* x_450; +x_436 = lean_ctor_get(x_435, 1); lean_inc(x_436); -x_437 = lean_ctor_get(x_435, 1); -lean_inc(x_437); lean_dec(x_435); -x_438 = lean_ctor_get(x_436, 0); +x_437 = lean_st_ref_take(x_3, x_436); +x_438 = lean_ctor_get(x_437, 0); lean_inc(x_438); -x_439 = lean_ctor_get(x_436, 1); +x_439 = lean_ctor_get(x_437, 1); lean_inc(x_439); -x_440 = lean_ctor_get(x_436, 2); +lean_dec(x_437); +x_440 = lean_ctor_get(x_438, 0); lean_inc(x_440); -if (lean_is_exclusive(x_436)) { - lean_ctor_release(x_436, 0); - lean_ctor_release(x_436, 1); - lean_ctor_release(x_436, 2); - x_441 = x_436; +x_441 = lean_ctor_get(x_438, 1); +lean_inc(x_441); +x_442 = lean_ctor_get(x_438, 2); +lean_inc(x_442); +if (lean_is_exclusive(x_438)) { + lean_ctor_release(x_438, 0); + lean_ctor_release(x_438, 1); + lean_ctor_release(x_438, 2); + x_443 = x_438; } else { - lean_dec_ref(x_436); - x_441 = lean_box(0); + lean_dec_ref(x_438); + x_443 = lean_box(0); } -x_442 = lean_array_push(x_439, x_429); -if (lean_is_scalar(x_441)) { - x_443 = lean_alloc_ctor(0, 3, 0); +x_444 = lean_array_push(x_441, x_430); +if (lean_is_scalar(x_443)) { + x_445 = lean_alloc_ctor(0, 3, 0); } else { - x_443 = x_441; + x_445 = x_443; } -lean_ctor_set(x_443, 0, x_438); -lean_ctor_set(x_443, 1, x_442); -lean_ctor_set(x_443, 2, x_440); -x_444 = lean_st_ref_set(x_3, x_443, x_437); +lean_ctor_set(x_445, 0, x_440); +lean_ctor_set(x_445, 1, x_444); +lean_ctor_set(x_445, 2, x_442); +x_446 = lean_st_ref_set(x_3, x_445, x_439); lean_dec(x_3); -x_445 = lean_ctor_get(x_444, 1); -lean_inc(x_445); -if (lean_is_exclusive(x_444)) { - lean_ctor_release(x_444, 0); - lean_ctor_release(x_444, 1); - x_446 = x_444; +x_447 = lean_ctor_get(x_446, 1); +lean_inc(x_447); +if (lean_is_exclusive(x_446)) { + lean_ctor_release(x_446, 0); + lean_ctor_release(x_446, 1); + x_448 = x_446; } else { - lean_dec_ref(x_444); - x_446 = lean_box(0); + lean_dec_ref(x_446); + x_448 = lean_box(0); } -x_447 = lean_box(0); -if (lean_is_scalar(x_446)) { - x_448 = lean_alloc_ctor(0, 2, 0); +x_449 = lean_box(0); +if (lean_is_scalar(x_448)) { + x_450 = lean_alloc_ctor(0, 2, 0); } else { - x_448 = x_446; + x_450 = x_448; } -lean_ctor_set(x_448, 0, x_447); -lean_ctor_set(x_448, 1, x_445); -return x_448; +lean_ctor_set(x_450, 0, x_449); +lean_ctor_set(x_450, 1, x_447); +return x_450; } else { -lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; -lean_dec(x_429); +lean_object* x_451; lean_object* x_452; lean_object* x_453; lean_object* x_454; +lean_dec(x_430); lean_dec(x_3); -x_449 = lean_ctor_get(x_433, 0); -lean_inc(x_449); -x_450 = lean_ctor_get(x_433, 1); -lean_inc(x_450); -if (lean_is_exclusive(x_433)) { - lean_ctor_release(x_433, 0); - lean_ctor_release(x_433, 1); - x_451 = x_433; +x_451 = lean_ctor_get(x_435, 0); +lean_inc(x_451); +x_452 = lean_ctor_get(x_435, 1); +lean_inc(x_452); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_453 = x_435; } else { - lean_dec_ref(x_433); - x_451 = lean_box(0); + lean_dec_ref(x_435); + x_453 = lean_box(0); } -if (lean_is_scalar(x_451)) { - x_452 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_453)) { + x_454 = lean_alloc_ctor(1, 2, 0); } else { - x_452 = x_451; + x_454 = x_453; } -lean_ctor_set(x_452, 0, x_449); -lean_ctor_set(x_452, 1, x_450); -return x_452; +lean_ctor_set(x_454, 0, x_451); +lean_ctor_set(x_454, 1, x_452); +return x_454; } } } else { -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; uint8_t x_460; +lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; uint8_t x_462; lean_dec(x_1); -x_453 = lean_ctor_get(x_364, 0); -lean_inc(x_453); -lean_dec(x_364); -x_454 = lean_ctor_get(x_363, 1); -lean_inc(x_454); -lean_dec(x_363); -x_455 = lean_ctor_get(x_2, 1); +x_455 = lean_ctor_get(x_365, 0); lean_inc(x_455); -x_456 = lean_ctor_get(x_453, 0); +lean_dec(x_365); +x_456 = lean_ctor_get(x_364, 1); lean_inc(x_456); -x_457 = lean_ctor_get(x_453, 1); +lean_dec(x_364); +x_457 = lean_ctor_get(x_2, 1); lean_inc(x_457); -x_458 = lean_ctor_get(x_453, 3); +x_458 = lean_ctor_get(x_455, 0); lean_inc(x_458); -lean_inc(x_456); -x_459 = lean_apply_1(x_455, x_456); -x_460 = lean_unbox(x_459); -lean_dec(x_459); -if (x_460 == 0) +x_459 = lean_ctor_get(x_455, 1); +lean_inc(x_459); +x_460 = lean_ctor_get(x_455, 3); +lean_inc(x_460); +lean_inc(x_458); +x_461 = lean_apply_1(x_457, x_458); +x_462 = lean_unbox(x_461); +lean_dec(x_461); +if (x_462 == 0) { -lean_object* x_461; +lean_object* x_463; +lean_dec(x_460); +lean_dec(x_459); lean_dec(x_458); -lean_dec(x_457); -lean_dec(x_456); lean_inc(x_3); -lean_inc(x_453); -x_461 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_453, x_2, x_3, x_4, x_5, x_6, x_7, x_454); -if (lean_obj_tag(x_461) == 0) +lean_inc(x_455); +x_463 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_455, x_2, x_3, x_4, x_5, x_6, x_7, x_456); +if (lean_obj_tag(x_463) == 0) { -lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; -x_462 = lean_ctor_get(x_461, 1); -lean_inc(x_462); -lean_dec(x_461); -x_463 = lean_st_ref_take(x_3, x_462); -x_464 = lean_ctor_get(x_463, 0); +lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; lean_object* x_478; lean_object* x_479; +x_464 = lean_ctor_get(x_463, 1); lean_inc(x_464); -x_465 = lean_ctor_get(x_463, 1); -lean_inc(x_465); lean_dec(x_463); -x_466 = lean_ctor_get(x_464, 0); +x_465 = lean_st_ref_take(x_3, x_464); +x_466 = lean_ctor_get(x_465, 0); lean_inc(x_466); -x_467 = lean_ctor_get(x_464, 1); +x_467 = lean_ctor_get(x_465, 1); lean_inc(x_467); -x_468 = lean_ctor_get(x_464, 2); +lean_dec(x_465); +x_468 = lean_ctor_get(x_466, 0); lean_inc(x_468); -if (lean_is_exclusive(x_464)) { - lean_ctor_release(x_464, 0); - lean_ctor_release(x_464, 1); - lean_ctor_release(x_464, 2); - x_469 = x_464; +x_469 = lean_ctor_get(x_466, 1); +lean_inc(x_469); +x_470 = lean_ctor_get(x_466, 2); +lean_inc(x_470); +if (lean_is_exclusive(x_466)) { + lean_ctor_release(x_466, 0); + lean_ctor_release(x_466, 1); + lean_ctor_release(x_466, 2); + x_471 = x_466; } else { - lean_dec_ref(x_464); - x_469 = lean_box(0); -} -x_470 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_470, 0, x_453); -x_471 = lean_array_push(x_468, x_470); -if (lean_is_scalar(x_469)) { - x_472 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_466); + x_471 = lean_box(0); +} +x_472 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_472, 0, x_455); +x_473 = lean_array_push(x_470, x_472); +if (lean_is_scalar(x_471)) { + x_474 = lean_alloc_ctor(0, 3, 0); } else { - x_472 = x_469; + x_474 = x_471; } -lean_ctor_set(x_472, 0, x_466); -lean_ctor_set(x_472, 1, x_467); -lean_ctor_set(x_472, 2, x_471); -x_473 = lean_st_ref_set(x_3, x_472, x_465); +lean_ctor_set(x_474, 0, x_468); +lean_ctor_set(x_474, 1, x_469); +lean_ctor_set(x_474, 2, x_473); +x_475 = lean_st_ref_set(x_3, x_474, x_467); lean_dec(x_3); -x_474 = lean_ctor_get(x_473, 1); -lean_inc(x_474); -if (lean_is_exclusive(x_473)) { - lean_ctor_release(x_473, 0); - lean_ctor_release(x_473, 1); - x_475 = x_473; +x_476 = lean_ctor_get(x_475, 1); +lean_inc(x_476); +if (lean_is_exclusive(x_475)) { + lean_ctor_release(x_475, 0); + lean_ctor_release(x_475, 1); + x_477 = x_475; } else { - lean_dec_ref(x_473); - x_475 = lean_box(0); + lean_dec_ref(x_475); + x_477 = lean_box(0); } -x_476 = lean_box(0); -if (lean_is_scalar(x_475)) { - x_477 = lean_alloc_ctor(0, 2, 0); +x_478 = lean_box(0); +if (lean_is_scalar(x_477)) { + x_479 = lean_alloc_ctor(0, 2, 0); } else { - x_477 = x_475; + x_479 = x_477; } -lean_ctor_set(x_477, 0, x_476); -lean_ctor_set(x_477, 1, x_474); -return x_477; +lean_ctor_set(x_479, 0, x_478); +lean_ctor_set(x_479, 1, x_476); +return x_479; } else { -lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; -lean_dec(x_453); +lean_object* x_480; lean_object* x_481; lean_object* x_482; lean_object* x_483; +lean_dec(x_455); lean_dec(x_3); -x_478 = lean_ctor_get(x_461, 0); -lean_inc(x_478); -x_479 = lean_ctor_get(x_461, 1); -lean_inc(x_479); -if (lean_is_exclusive(x_461)) { - lean_ctor_release(x_461, 0); - lean_ctor_release(x_461, 1); - x_480 = x_461; +x_480 = lean_ctor_get(x_463, 0); +lean_inc(x_480); +x_481 = lean_ctor_get(x_463, 1); +lean_inc(x_481); +if (lean_is_exclusive(x_463)) { + lean_ctor_release(x_463, 0); + lean_ctor_release(x_463, 1); + x_482 = x_463; } else { - lean_dec_ref(x_461); - x_480 = lean_box(0); + lean_dec_ref(x_463); + x_482 = lean_box(0); } -if (lean_is_scalar(x_480)) { - x_481 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_482)) { + x_483 = lean_alloc_ctor(1, 2, 0); } else { - x_481 = x_480; + x_483 = x_482; } -lean_ctor_set(x_481, 0, x_478); -lean_ctor_set(x_481, 1, x_479); -return x_481; +lean_ctor_set(x_483, 0, x_480); +lean_ctor_set(x_483, 1, x_481); +return x_483; } } else { -lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; uint8_t x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; -lean_dec(x_453); +lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; uint8_t x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; +lean_dec(x_455); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_482 = lean_st_ref_take(x_3, x_454); -x_483 = lean_ctor_get(x_482, 0); -lean_inc(x_483); -x_484 = lean_ctor_get(x_482, 1); -lean_inc(x_484); -lean_dec(x_482); -x_485 = lean_ctor_get(x_483, 0); +x_484 = lean_st_ref_take(x_3, x_456); +x_485 = lean_ctor_get(x_484, 0); lean_inc(x_485); -x_486 = lean_ctor_get(x_483, 1); +x_486 = lean_ctor_get(x_484, 1); lean_inc(x_486); -x_487 = lean_ctor_get(x_483, 2); +lean_dec(x_484); +x_487 = lean_ctor_get(x_485, 0); lean_inc(x_487); -if (lean_is_exclusive(x_483)) { - lean_ctor_release(x_483, 0); - lean_ctor_release(x_483, 1); - lean_ctor_release(x_483, 2); - x_488 = x_483; +x_488 = lean_ctor_get(x_485, 1); +lean_inc(x_488); +x_489 = lean_ctor_get(x_485, 2); +lean_inc(x_489); +if (lean_is_exclusive(x_485)) { + lean_ctor_release(x_485, 0); + lean_ctor_release(x_485, 1); + lean_ctor_release(x_485, 2); + x_490 = x_485; } else { - lean_dec_ref(x_483); - x_488 = lean_box(0); -} -x_489 = 0; -x_490 = lean_alloc_ctor(0, 3, 1); -lean_ctor_set(x_490, 0, x_456); -lean_ctor_set(x_490, 1, x_457); -lean_ctor_set(x_490, 2, x_458); -lean_ctor_set_uint8(x_490, sizeof(void*)*3, x_489); -x_491 = lean_array_push(x_486, x_490); -if (lean_is_scalar(x_488)) { - x_492 = lean_alloc_ctor(0, 3, 0); + lean_dec_ref(x_485); + x_490 = lean_box(0); +} +x_491 = 0; +x_492 = lean_alloc_ctor(0, 3, 1); +lean_ctor_set(x_492, 0, x_458); +lean_ctor_set(x_492, 1, x_459); +lean_ctor_set(x_492, 2, x_460); +lean_ctor_set_uint8(x_492, sizeof(void*)*3, x_491); +x_493 = lean_array_push(x_488, x_492); +if (lean_is_scalar(x_490)) { + x_494 = lean_alloc_ctor(0, 3, 0); } else { - x_492 = x_488; + x_494 = x_490; } -lean_ctor_set(x_492, 0, x_485); -lean_ctor_set(x_492, 1, x_491); -lean_ctor_set(x_492, 2, x_487); -x_493 = lean_st_ref_set(x_3, x_492, x_484); +lean_ctor_set(x_494, 0, x_487); +lean_ctor_set(x_494, 1, x_493); +lean_ctor_set(x_494, 2, x_489); +x_495 = lean_st_ref_set(x_3, x_494, x_486); lean_dec(x_3); -x_494 = lean_ctor_get(x_493, 1); -lean_inc(x_494); -if (lean_is_exclusive(x_493)) { - lean_ctor_release(x_493, 0); - lean_ctor_release(x_493, 1); - x_495 = x_493; +x_496 = lean_ctor_get(x_495, 1); +lean_inc(x_496); +if (lean_is_exclusive(x_495)) { + lean_ctor_release(x_495, 0); + lean_ctor_release(x_495, 1); + x_497 = x_495; } else { - lean_dec_ref(x_493); - x_495 = lean_box(0); + lean_dec_ref(x_495); + x_497 = lean_box(0); } -x_496 = lean_box(0); -if (lean_is_scalar(x_495)) { - x_497 = lean_alloc_ctor(0, 2, 0); +x_498 = lean_box(0); +if (lean_is_scalar(x_497)) { + x_499 = lean_alloc_ctor(0, 2, 0); } else { - x_497 = x_495; + x_499 = x_497; } -lean_ctor_set(x_497, 0, x_496); -lean_ctor_set(x_497, 1, x_494); -return x_497; +lean_ctor_set(x_499, 0, x_498); +lean_ctor_set(x_499, 1, x_496); +return x_499; } } } } else { -lean_object* x_498; lean_object* x_499; -lean_dec(x_354); +lean_object* x_500; lean_object* x_501; +lean_dec(x_355); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -3376,11 +3421,11 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_498 = lean_box(0); -x_499 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_499, 0, x_498); -lean_ctor_set(x_499, 1, x_352); -return x_499; +x_500 = lean_box(0); +x_501 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_501, 0, x_500); +lean_ctor_set(x_501, 1, x_353); +return x_501; } } } @@ -3751,14 +3796,15 @@ return x_12; } default: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; x_13 = lean_ctor_get(x_1, 0); lean_inc(x_13); lean_dec(x_1); x_14 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_15 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); -x_16 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_14, x_15, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_16; +x_16 = 0; +x_17 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_14, x_15, x_13, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_17; } } } @@ -3766,49 +3812,50 @@ return x_16; lean_object* l_Lean_Compiler_LCNF_Closure_collectFunDecl(lean_object* x_1, lean_object* x_2, 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_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; x_9 = lean_ctor_get(x_1, 3); lean_inc(x_9); x_10 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_11 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_12 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_12 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_10, x_11, x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_12) == 0) +x_13 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_10, x_11, x_9, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_ctor_get(x_1, 2); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_ctor_get(x_1, 2); +lean_inc(x_15); 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_15 = l_Lean_Compiler_LCNF_Closure_collectParams(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); +x_16 = l_Lean_Compiler_LCNF_Closure_collectParams(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_14); lean_dec(x_15); -x_17 = lean_ctor_get(x_1, 4); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); +lean_dec(x_16); +x_18 = lean_ctor_get(x_1, 4); +lean_inc(x_18); lean_dec(x_1); -x_18 = l_Lean_Compiler_LCNF_Closure_collectCode(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_16); -return x_18; +x_19 = l_Lean_Compiler_LCNF_Closure_collectCode(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +return x_19; } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -3816,29 +3863,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_19 = !lean_is_exclusive(x_15); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_16); +if (x_20 == 0) { -return x_15; +return x_16; } else { -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_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_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; +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; } } } else { -uint8_t x_23; +uint8_t x_24; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -3846,23 +3893,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_23 = !lean_is_exclusive(x_12); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_13); +if (x_24 == 0) { -return x_12; +return x_13; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -x_25 = lean_ctor_get(x_12, 1); +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_inc(x_24); -lean_dec(x_12); -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_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; } } } @@ -4060,7 +4107,7 @@ lean_object* l_Lean_Compiler_LCNF_Closure_collectCode(lean_object* x_1, lean_obj switch (lean_obj_tag(x_1)) { case 0: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; x_9 = lean_ctor_get(x_1, 0); lean_inc(x_9); x_10 = lean_ctor_get(x_1, 1); @@ -4070,21 +4117,22 @@ x_11 = lean_ctor_get(x_9, 2); lean_inc(x_11); x_12 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); x_13 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_14 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_14 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_12, x_13, x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_14) == 0) +x_15 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_12, x_13, x_11, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_ctor_get(x_9, 3); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_ctor_get(x_9, 3); +lean_inc(x_17); lean_dec(x_9); lean_inc(x_7); lean_inc(x_6); @@ -4092,20 +4140,20 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_17 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_Compiler_LCNF_Closure_collectLetValue(x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_16); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); x_1 = x_10; -x_8 = x_18; +x_8 = x_19; goto _start; } else { -uint8_t x_20; +uint8_t x_21; lean_dec(x_10); lean_dec(x_7); lean_dec(x_6); @@ -4113,29 +4161,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) { -return x_17; +return x_18; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_17, 0); -x_22 = lean_ctor_get(x_17, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_18, 0); +x_23 = lean_ctor_get(x_18, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_17); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_dec(x_18); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } } else { -uint8_t x_24; +uint8_t x_25; lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); @@ -4144,333 +4192,335 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_24 = !lean_is_exclusive(x_14); -if (x_24 == 0) +x_25 = !lean_is_exclusive(x_15); +if (x_25 == 0) { -return x_14; +return x_15; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_14, 0); -x_26 = lean_ctor_get(x_14, 1); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_15, 0); +x_27 = lean_ctor_get(x_15, 1); +lean_inc(x_27); lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_14); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_dec(x_15); +x_28 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } } case 3: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; -x_28 = lean_ctor_get(x_1, 1); -lean_inc(x_28); +lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_29 = lean_ctor_get(x_1, 1); +lean_inc(x_29); lean_dec(x_1); -x_29 = lean_array_get_size(x_28); -x_30 = lean_unsigned_to_nat(0u); -x_31 = lean_nat_dec_lt(x_30, x_29); -if (x_31 == 0) +x_30 = lean_array_get_size(x_29); +x_31 = lean_unsigned_to_nat(0u); +x_32 = lean_nat_dec_lt(x_31, x_30); +if (x_32 == 0) { -lean_object* x_32; lean_object* x_33; +lean_object* x_33; lean_object* x_34; +lean_dec(x_30); lean_dec(x_29); -lean_dec(x_28); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_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_8); -return x_33; +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_8); +return x_34; } else { -uint8_t x_34; -x_34 = lean_nat_dec_le(x_29, x_29); -if (x_34 == 0) +uint8_t x_35; +x_35 = lean_nat_dec_le(x_30, x_30); +if (x_35 == 0) { -lean_object* x_35; lean_object* x_36; +lean_object* x_36; lean_object* x_37; +lean_dec(x_30); lean_dec(x_29); -lean_dec(x_28); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_8); -return x_36; +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_8); +return x_37; } else { -size_t x_37; size_t x_38; lean_object* x_39; lean_object* x_40; -x_37 = 0; -x_38 = lean_usize_of_nat(x_29); +size_t x_38; size_t x_39; lean_object* x_40; lean_object* x_41; +x_38 = 0; +x_39 = lean_usize_of_nat(x_30); +lean_dec(x_30); +x_40 = lean_box(0); +x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectLetValue___spec__1(x_29, x_38, x_39, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_29); -x_39 = lean_box(0); -x_40 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectLetValue___spec__1(x_28, x_37, x_38, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_28); -return x_40; +return x_41; } } } case 4: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_41 = lean_ctor_get(x_1, 0); -lean_inc(x_41); -lean_dec(x_1); -x_42 = lean_ctor_get(x_41, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; +x_42 = lean_ctor_get(x_1, 0); lean_inc(x_42); -x_43 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); -x_44 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +lean_dec(x_1); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +x_44 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); +x_45 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_46 = 0; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_45 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_43, x_44, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_45) == 0) -{ -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 = lean_ctor_get(x_41, 2); -lean_inc(x_47); +x_47 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_44, x_45, x_43, x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_47) == 0) +{ +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 = lean_ctor_get(x_42, 2); +lean_inc(x_49); 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_48 = l_Lean_Compiler_LCNF_Closure_collectFVar(x_47, x_2, x_3, x_4, x_5, x_6, x_7, x_46); -if (lean_obj_tag(x_48) == 0) -{ -uint8_t x_49; -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +x_50 = l_Lean_Compiler_LCNF_Closure_collectFVar(x_49, x_2, x_3, x_4, x_5, x_6, x_7, x_48); +if (lean_obj_tag(x_50) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_50 = lean_ctor_get(x_48, 1); -x_51 = lean_ctor_get(x_48, 0); -lean_dec(x_51); -x_52 = lean_ctor_get(x_41, 3); -lean_inc(x_52); -lean_dec(x_41); -x_53 = lean_array_get_size(x_52); -x_54 = lean_unsigned_to_nat(0u); -x_55 = lean_nat_dec_lt(x_54, x_53); -if (x_55 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_56; +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_52 = lean_ctor_get(x_50, 1); +x_53 = lean_ctor_get(x_50, 0); lean_dec(x_53); -lean_dec(x_52); +x_54 = lean_ctor_get(x_42, 3); +lean_inc(x_54); +lean_dec(x_42); +x_55 = lean_array_get_size(x_54); +x_56 = lean_unsigned_to_nat(0u); +x_57 = lean_nat_dec_lt(x_56, x_55); +if (x_57 == 0) +{ +lean_object* x_58; +lean_dec(x_55); +lean_dec(x_54); 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_56 = lean_box(0); -lean_ctor_set(x_48, 0, x_56); -return x_48; +x_58 = lean_box(0); +lean_ctor_set(x_50, 0, x_58); +return x_50; } else { -uint8_t x_57; -x_57 = lean_nat_dec_le(x_53, x_53); -if (x_57 == 0) +uint8_t x_59; +x_59 = lean_nat_dec_le(x_55, x_55); +if (x_59 == 0) { -lean_object* x_58; -lean_dec(x_53); -lean_dec(x_52); +lean_object* x_60; +lean_dec(x_55); +lean_dec(x_54); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_58 = lean_box(0); -lean_ctor_set(x_48, 0, x_58); -return x_48; +x_60 = lean_box(0); +lean_ctor_set(x_50, 0, x_60); +return x_50; } else { -size_t x_59; size_t x_60; lean_object* x_61; lean_object* x_62; -lean_free_object(x_48); -x_59 = 0; -x_60 = lean_usize_of_nat(x_53); -lean_dec(x_53); -x_61 = lean_box(0); -x_62 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___spec__1(x_52, x_59, x_60, x_61, x_2, x_3, x_4, x_5, x_6, x_7, x_50); -lean_dec(x_52); -return x_62; +size_t x_61; size_t x_62; lean_object* x_63; lean_object* x_64; +lean_free_object(x_50); +x_61 = 0; +x_62 = lean_usize_of_nat(x_55); +lean_dec(x_55); +x_63 = lean_box(0); +x_64 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___spec__1(x_54, x_61, x_62, x_63, x_2, x_3, x_4, x_5, x_6, x_7, x_52); +lean_dec(x_54); +return x_64; } } } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_63 = lean_ctor_get(x_48, 1); -lean_inc(x_63); -lean_dec(x_48); -x_64 = lean_ctor_get(x_41, 3); -lean_inc(x_64); -lean_dec(x_41); -x_65 = lean_array_get_size(x_64); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; -lean_dec(x_65); -lean_dec(x_64); +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_65 = lean_ctor_get(x_50, 1); +lean_inc(x_65); +lean_dec(x_50); +x_66 = lean_ctor_get(x_42, 3); +lean_inc(x_66); +lean_dec(x_42); +x_67 = lean_array_get_size(x_66); +x_68 = lean_unsigned_to_nat(0u); +x_69 = lean_nat_dec_lt(x_68, x_67); +if (x_69 == 0) +{ +lean_object* x_70; lean_object* x_71; +lean_dec(x_67); +lean_dec(x_66); 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_68 = lean_box(0); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_63); -return x_69; +x_70 = lean_box(0); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_65); +return x_71; } else { -uint8_t x_70; -x_70 = lean_nat_dec_le(x_65, x_65); -if (x_70 == 0) +uint8_t x_72; +x_72 = lean_nat_dec_le(x_67, x_67); +if (x_72 == 0) { -lean_object* x_71; lean_object* x_72; -lean_dec(x_65); -lean_dec(x_64); +lean_object* x_73; lean_object* x_74; +lean_dec(x_67); +lean_dec(x_66); 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_71 = lean_box(0); -x_72 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_63); -return x_72; +x_73 = lean_box(0); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_65); +return x_74; } else { -size_t x_73; size_t x_74; lean_object* x_75; lean_object* x_76; -x_73 = 0; -x_74 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_75 = lean_box(0); -x_76 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___spec__1(x_64, x_73, x_74, x_75, x_2, x_3, x_4, x_5, x_6, x_7, x_63); -lean_dec(x_64); -return x_76; +size_t x_75; size_t x_76; lean_object* x_77; lean_object* x_78; +x_75 = 0; +x_76 = lean_usize_of_nat(x_67); +lean_dec(x_67); +x_77 = lean_box(0); +x_78 = l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectCode___spec__1(x_66, x_75, x_76, x_77, x_2, x_3, x_4, x_5, x_6, x_7, x_65); +lean_dec(x_66); +return x_78; } } } } else { -uint8_t x_77; -lean_dec(x_41); +uint8_t x_79; +lean_dec(x_42); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_77 = !lean_is_exclusive(x_48); -if (x_77 == 0) +x_79 = !lean_is_exclusive(x_50); +if (x_79 == 0) { -return x_48; +return x_50; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_48, 0); -x_79 = lean_ctor_get(x_48, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_48); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_50, 0); +x_81 = lean_ctor_get(x_50, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_50); +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; } } } else { -uint8_t x_81; -lean_dec(x_41); +uint8_t x_83; +lean_dec(x_42); 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_81 = !lean_is_exclusive(x_45); -if (x_81 == 0) +x_83 = !lean_is_exclusive(x_47); +if (x_83 == 0) { -return x_45; +return x_47; } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; -x_82 = lean_ctor_get(x_45, 0); -x_83 = lean_ctor_get(x_45, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_45); -x_84 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_84, 0, x_82); -lean_ctor_set(x_84, 1, x_83); -return x_84; +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_47, 0); +x_85 = lean_ctor_get(x_47, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_47); +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; } } } case 5: { -lean_object* x_85; lean_object* x_86; -x_85 = lean_ctor_get(x_1, 0); -lean_inc(x_85); +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_1, 0); +lean_inc(x_87); lean_dec(x_1); -x_86 = l_Lean_Compiler_LCNF_Closure_collectFVar(x_85, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_86; +x_88 = l_Lean_Compiler_LCNF_Closure_collectFVar(x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_88; } case 6: { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_87 = lean_ctor_get(x_1, 0); -lean_inc(x_87); +lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; +x_89 = lean_ctor_get(x_1, 0); +lean_inc(x_89); lean_dec(x_1); -x_88 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); -x_89 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); -x_90 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_88, x_89, x_87, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_90; +x_90 = lean_alloc_closure((void*)(l_Lean_Expr_isFVar___boxed), 1, 0); +x_91 = lean_alloc_closure((void*)(l_Lean_Compiler_LCNF_Closure_collectType___lambda__1), 8, 0); +x_92 = 0; +x_93 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_90, x_91, x_89, x_92, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_93; } default: { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_1, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_1, 1); -lean_inc(x_92); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_1, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_1, 1); +lean_inc(x_95); lean_dec(x_1); lean_inc(x_7); lean_inc(x_6); @@ -4478,44 +4528,44 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_93 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_91, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_93) == 0) +x_96 = l_Lean_Compiler_LCNF_Closure_collectFunDecl(x_94, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_94; -x_94 = lean_ctor_get(x_93, 1); -lean_inc(x_94); -lean_dec(x_93); -x_1 = x_92; -x_8 = x_94; +lean_object* x_97; +x_97 = lean_ctor_get(x_96, 1); +lean_inc(x_97); +lean_dec(x_96); +x_1 = x_95; +x_8 = x_97; goto _start; } else { -uint8_t x_96; -lean_dec(x_92); +uint8_t x_99; +lean_dec(x_95); 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_96 = !lean_is_exclusive(x_93); -if (x_96 == 0) +x_99 = !lean_is_exclusive(x_96); +if (x_99 == 0) { -return x_93; +return x_96; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_93, 0); -x_98 = lean_ctor_get(x_93, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_93); -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_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_96, 0); +x_101 = lean_ctor_get(x_96, 1); +lean_inc(x_101); +lean_inc(x_100); +lean_dec(x_96); +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; } } } @@ -4574,6 +4624,36 @@ lean_dec(x_2); return x_10; } } +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; lean_object* x_15; +x_14 = lean_unbox(x_4); +lean_dec(x_4); +x_15 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__2___lambda__1(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_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___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: +{ +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_3); +lean_dec(x_3); +x_14 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Compiler_LCNF_Closure_collectType___spec__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; +} +} +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_4); +lean_dec(x_4); +x_13 = l_Lean_ForEachExprWhere_visit___at_Lean_Compiler_LCNF_Closure_collectType___spec__1(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +} LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_Closure_collectLetValue___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: { diff --git a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c index 1651164811e3..322bca267544 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/FloatLetIn.c @@ -8518,7 +8518,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(212u); +x_3 = lean_unsigned_to_nat(213u); 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 fd6f3bdfebc2..29d104379816 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c @@ -8464,7 +8464,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(212u); +x_3 = lean_unsigned_to_nat(213u); 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/Compiler/LCNF/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Main.c index c2c9cb68bff0..0ae7669e4e87 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -16,19 +16,20 @@ extern "C" { 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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__1; 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*); -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_PassManager_run___spec__6___lambda__1___closed__9; 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*); +static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__3; 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; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_PassManager_run___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Compiler_LCNF_compile___closed__1; -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*); @@ -37,11 +38,12 @@ lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_o double lean_float_div(double, double); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__28; lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__12; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___closed__4; +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, 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*, 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*, 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_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*); @@ -49,8 +51,9 @@ static lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run 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_PassManager_run___spec__6___lambda__1___closed__18; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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*); @@ -63,6 +66,7 @@ extern lean_object* l_Lean_Expr_instBEq; extern lean_object* l_Lean_casesOnSuffix; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__20; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(lean_object*, lean_object*, lean_object*, 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*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__21; @@ -71,145 +75,154 @@ lean_object* l_Lean_Compiler_LCNF_getPassManager___rarg(lean_object*, lean_objec 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; +lean_object* lean_io_get_num_heartbeats(lean_object*); 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_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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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_lcnf_compile_decls(lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__24; lean_object* l_Lean_Compiler_LCNF_compile(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___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* 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*, lean_object*, uint8_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___lambda__2___closed__13; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__20; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__4; +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__16; 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*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__7___boxed(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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__7; 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__8; 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*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3(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_Lean_Compiler_LCNF_shouldGenerateCode___closed__11; 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*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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 double l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__5; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_PassManager_run___spec__9(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__23; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; 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_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__6___lambda__1___closed__7; 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__5; 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; 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; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__10; lean_object* lean_io_mono_nanos_now(lean_object*); -static double 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_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__31; -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*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); lean_object* l_instHashableProd___rarg___boxed(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; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___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_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__13; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; lean_object* l_Lean_Compiler_LCNF_showDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isMatcher___at_Lean_Compiler_LCNF_shouldGenerateCode___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isExtern(lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_threshold; 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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__3___closed__8; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__16; static lean_object* l_Lean_Compiler_LCNF_compile___closed__2; +extern lean_object* l_Lean_diagnostics; +LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_PassManager_run___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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; -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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__14; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__5(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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__10; 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; 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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__2; 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_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__7___lambda__1___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__1; 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_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4(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__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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___closed__2; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__6; static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_CompilerM_run___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__22; -static double l_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; 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_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); 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___closed__3; +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*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__22; -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*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__3; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__6; lean_object* l_Lean_Compiler_LCNF_CompilerM_run___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__1(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_Compiler_LCNF_shouldGenerateCode___closed__4; extern lean_object* l_Lean_Expr_instHashable; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___closed__14; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Compiler_LCNF_PassManager_run___spec__3(lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__4; 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; -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; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__10; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1(lean_object*, 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__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*, 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*); +static double l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getDeclInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_checkpoint(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -221,19 +234,20 @@ 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__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*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(lean_object*, uint8_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_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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__19; 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*); +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*, 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__17; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__16; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__11; uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); @@ -241,40 +255,39 @@ lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); 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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__3; 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__15; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__4; 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*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__6; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_mapTR_loop___at_Lean_compileDecl___spec__1(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__6___lambda__1___closed__11; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__5; static lean_object* l_Lean_Compiler_LCNF_compile___closed__3; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__18; 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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___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*); -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_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*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(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_Lean_Compiler_LCNF_shouldGenerateCode___lambda__6(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__2___closed__2; 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; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__2; 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__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); 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) { @@ -1322,559 +1335,666 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_check _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Compiler", 8); +x_1 = l_Lean_maxRecDepth; return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("size: ", 6); +return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("pp", 2); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("motives", 7); +x_1 = lean_mk_string_from_bytes("\n", 1); return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("pi", 2); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__4; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__3; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__4; -x_3 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; } } static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_box(0); -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__8() { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___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_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +lean_dec(x_7); +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; uint8_t x_18; +x_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; +x_15 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__1(x_1, x_14); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); +lean_inc(x_3); +x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(x_3, x_4, x_5, x_8, x_9, x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; +lean_dec(x_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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_6, x_20, x_4, x_5, x_8, x_9, x_19); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_16, 1); +lean_inc(x_22); +lean_dec(x_16); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_6); +x_23 = l_Lean_Compiler_LCNF_ppDecl_x27(x_6, x_8, x_9, x_22); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +lean_inc(x_6); +x_26 = l_Lean_Compiler_LCNF_Decl_size(x_6); +x_27 = l___private_Init_Data_Repr_0__Nat_reprFast(x_26); +x_28 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_28, 0, x_27); +x_29 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__3; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_34, 0, x_24); +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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; +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_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(x_3, x_37, x_4, x_5, x_8, x_9, x_25); +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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_6, x_39, x_4, x_5, x_8, x_9, x_40); +lean_dec(x_39); +return x_41; +} +else +{ +uint8_t x_42; +lean_dec(x_8); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_42 = !lean_is_exclusive(x_23); +if (x_42 == 0) +{ +return x_23; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_23, 0); +x_44 = lean_ctor_get(x_23, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_23); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +} +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; 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_object* x_60; uint8_t x_61; +x_46 = lean_ctor_get(x_8, 0); +x_47 = lean_ctor_get(x_8, 1); +x_48 = lean_ctor_get(x_8, 3); +x_49 = lean_ctor_get(x_8, 5); +x_50 = lean_ctor_get(x_8, 6); +x_51 = lean_ctor_get(x_8, 7); +x_52 = lean_ctor_get(x_8, 8); +x_53 = lean_ctor_get(x_8, 9); +x_54 = lean_ctor_get(x_8, 10); +x_55 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_8); +x_56 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; +x_57 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__1(x_1, x_56); +x_58 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_58, 0, x_46); +lean_ctor_set(x_58, 1, x_47); +lean_ctor_set(x_58, 2, x_1); +lean_ctor_set(x_58, 3, x_48); +lean_ctor_set(x_58, 4, x_57); +lean_ctor_set(x_58, 5, x_49); +lean_ctor_set(x_58, 6, x_50); +lean_ctor_set(x_58, 7, x_51); +lean_ctor_set(x_58, 8, x_52); +lean_ctor_set(x_58, 9, x_53); +lean_ctor_set(x_58, 10, x_54); +lean_ctor_set_uint8(x_58, sizeof(void*)*11, x_55); +lean_ctor_set_uint8(x_58, sizeof(void*)*11 + 1, x_2); +lean_inc(x_3); +x_59 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(x_3, x_4, x_5, x_58, x_9, x_10); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_unbox(x_60); +lean_dec(x_60); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_dec(x_3); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_box(0); +x_64 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_6, x_63, x_4, x_5, x_58, x_9, x_62); +return x_64; +} +else +{ +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_59, 1); +lean_inc(x_65); +lean_dec(x_59); +lean_inc(x_9); +lean_inc(x_58); +lean_inc(x_6); +x_66 = l_Lean_Compiler_LCNF_ppDecl_x27(x_6, x_58, x_9, 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; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_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); +lean_inc(x_6); +x_69 = l_Lean_Compiler_LCNF_Decl_size(x_6); +x_70 = l___private_Init_Data_Repr_0__Nat_reprFast(x_69); +x_71 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_71, 0, x_70); +x_72 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_72, 0, x_71); +x_73 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__3; +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_72); +x_75 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__5; +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_77, 0, x_67); +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(x_3, x_80, x_4, x_5, x_58, x_9, x_68); +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_6, x_82, x_4, x_5, x_58, x_9, x_83); +lean_dec(x_82); +return x_84; +} +else +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +lean_dec(x_58); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_85 = lean_ctor_get(x_66, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_66, 1); +lean_inc(x_86); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + lean_ctor_release(x_66, 1); + x_87 = x_66; +} else { + lean_dec_ref(x_66); + x_87 = lean_box(0); +} +if (lean_is_scalar(x_87)) { + x_88 = lean_alloc_ctor(1, 2, 0); +} else { + x_88 = x_87; +} +lean_ctor_set(x_88, 0, x_85); +lean_ctor_set(x_88, 1, x_86); +return x_88; +} +} +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("size: ", 6); +x_1 = lean_mk_string_from_bytes("Compiler", 8); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__8; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\n", 1); +x_1 = lean_mk_string_from_bytes("pp", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__10; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("motives", 7); +return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__12() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 0); +x_1 = lean_mk_string_from_bytes("pi", 2); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__4; +x_3 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__5; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__12; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_Compiler_LCNF_shouldGenerateCode___closed__13; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9() { _start: { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__2; -x_10 = l_Lean_Name_append(x_9, x_1); -x_11 = !lean_is_exclusive(x_6); -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; uint8_t x_18; -x_12 = lean_ctor_get(x_6, 2); -x_13 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6; -x_14 = 0; -x_15 = l_Lean_KVMap_setBool(x_12, x_13, x_14); -lean_ctor_set(x_6, 2, x_15); -lean_inc(x_10); -x_16 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(x_10, x_4, x_5, x_6, x_7, x_8); -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -lean_dec(x_10); -x_19 = lean_ctor_get(x_16, 1); +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; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; +x_9 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__2; +x_10 = l_Lean_Name_append(x_9, x_1); +x_11 = lean_ctor_get(x_6, 2); +lean_inc(x_11); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__6; +x_13 = 0; +x_14 = l_Lean_KVMap_setBool(x_11, x_12, x_13); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__7; +x_16 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(x_14, x_15); +x_17 = lean_st_ref_get(x_7, x_8); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); -lean_dec(x_16); -x_20 = lean_box(0); -x_21 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_2, x_20, x_4, x_5, x_6, x_7, x_19); -if (lean_obj_tag(x_21) == 0) +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +lean_dec(x_20); +if (x_21 == 0) { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) +if (x_16 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_21, 0); -lean_dec(x_23); -x_24 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; -lean_ctor_set(x_21, 0, x_24); -return x_21; +uint8_t x_79; +x_79 = 1; +x_22 = x_79; +goto block_78; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_21, 1); -lean_inc(x_25); -lean_dec(x_21); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__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; +x_22 = x_13; +goto block_78; } } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_21); -if (x_28 == 0) +if (x_16 == 0) { -return x_21; +x_22 = x_13; +goto block_78; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_21, 0); -x_30 = lean_ctor_get(x_21, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_21); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; -} +uint8_t x_80; +x_80 = 1; +x_22 = x_80; +goto block_78; } } -else -{ -lean_object* x_32; lean_object* x_33; -x_32 = lean_ctor_get(x_16, 1); -lean_inc(x_32); -lean_dec(x_16); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_2); -x_33 = l_Lean_Compiler_LCNF_ppDecl_x27(x_2, x_6, x_7, x_32); -if (lean_obj_tag(x_33) == 0) +block_78: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; 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_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -lean_inc(x_2); -x_36 = l_Lean_Compiler_LCNF_Decl_size(x_2); -x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_36); -x_38 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_38, 0, x_37); -x_39 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_39, 0, x_38); -x_40 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__9; -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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__11; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_41); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_44, 0, x_34); -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_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(x_10, x_47, x_4, x_5, x_6, x_7, x_35); -x_49 = lean_ctor_get(x_48, 0); -lean_inc(x_49); -x_50 = lean_ctor_get(x_48, 1); -lean_inc(x_50); -lean_dec(x_48); -x_51 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_2, x_49, x_4, x_5, x_6, x_7, x_50); -lean_dec(x_49); -if (lean_obj_tag(x_51) == 0) -{ -uint8_t x_52; -x_52 = !lean_is_exclusive(x_51); -if (x_52 == 0) +if (x_22 == 0) { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_51, 0); -lean_dec(x_53); -x_54 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; -lean_ctor_set(x_51, 0, x_54); -return x_51; -} -else +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_st_ref_take(x_7, x_19); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_51, 1); -lean_inc(x_55); -lean_dec(x_51); -x_56 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_55); -return x_57; -} -} -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; +x_27 = lean_ctor_get(x_24, 0); +x_28 = lean_ctor_get(x_24, 4); +lean_dec(x_28); +x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__8; +lean_ctor_set(x_24, 4, x_30); +lean_ctor_set(x_24, 0, x_29); +x_31 = lean_st_ref_set(x_7, x_24, x_25); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_14, x_16, x_10, x_4, x_5, x_2, x_33, x_6, x_7, x_32); +if (lean_obj_tag(x_34) == 0) { -uint8_t x_58; -x_58 = !lean_is_exclusive(x_51); -if (x_58 == 0) +uint8_t x_35; +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) { -return x_51; +lean_object* x_36; lean_object* x_37; +x_36 = lean_ctor_get(x_34, 0); +lean_dec(x_36); +x_37 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; +lean_ctor_set(x_34, 0, x_37); +return x_34; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_51, 0); -x_60 = lean_ctor_get(x_51, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_51); -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_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 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +return x_40; } } else { -uint8_t x_62; -lean_dec(x_6); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_62 = !lean_is_exclusive(x_33); -if (x_62 == 0) +uint8_t x_41; +x_41 = !lean_is_exclusive(x_34); +if (x_41 == 0) { -return x_33; +return x_34; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_33, 0); -x_64 = lean_ctor_get(x_33, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_33); -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_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_34, 0); +x_43 = lean_ctor_get(x_34, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_34); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; uint8_t x_78; lean_object* x_79; uint8_t x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_66 = lean_ctor_get(x_6, 0); -x_67 = lean_ctor_get(x_6, 1); -x_68 = lean_ctor_get(x_6, 2); -x_69 = lean_ctor_get(x_6, 3); -x_70 = lean_ctor_get(x_6, 4); -x_71 = lean_ctor_get(x_6, 5); -x_72 = lean_ctor_get(x_6, 6); -x_73 = lean_ctor_get(x_6, 7); -x_74 = lean_ctor_get(x_6, 8); -x_75 = lean_ctor_get(x_6, 9); -x_76 = lean_ctor_get(x_6, 10); -x_77 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); -x_78 = lean_ctor_get_uint8(x_6, sizeof(void*)*11 + 1); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_inc(x_72); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_6); -x_79 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6; -x_80 = 0; -x_81 = l_Lean_KVMap_setBool(x_68, x_79, x_80); -x_82 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_82, 0, x_66); -lean_ctor_set(x_82, 1, x_67); -lean_ctor_set(x_82, 2, x_81); -lean_ctor_set(x_82, 3, x_69); -lean_ctor_set(x_82, 4, x_70); -lean_ctor_set(x_82, 5, x_71); -lean_ctor_set(x_82, 6, x_72); -lean_ctor_set(x_82, 7, x_73); -lean_ctor_set(x_82, 8, x_74); -lean_ctor_set(x_82, 9, x_75); -lean_ctor_set(x_82, 10, x_76); -lean_ctor_set_uint8(x_82, sizeof(void*)*11, x_77); -lean_ctor_set_uint8(x_82, sizeof(void*)*11 + 1, x_78); -lean_inc(x_10); -x_83 = l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__1(x_10, x_4, x_5, x_82, x_7, x_8); -x_84 = lean_ctor_get(x_83, 0); -lean_inc(x_84); -x_85 = lean_unbox(x_84); -lean_dec(x_84); -if (x_85 == 0) -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_10); -x_86 = lean_ctor_get(x_83, 1); -lean_inc(x_86); -lean_dec(x_83); -x_87 = lean_box(0); -x_88 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_2, x_87, x_4, x_5, x_82, x_7, x_86); -if (lean_obj_tag(x_88) == 0) +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_45 = lean_ctor_get(x_24, 0); +x_46 = lean_ctor_get(x_24, 1); +x_47 = lean_ctor_get(x_24, 2); +x_48 = lean_ctor_get(x_24, 3); +x_49 = lean_ctor_get(x_24, 5); +x_50 = lean_ctor_get(x_24, 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_24); +x_51 = l_Lean_Kernel_enableDiag(x_45, x_16); +x_52 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__8; +x_53 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_46); +lean_ctor_set(x_53, 2, x_47); +lean_ctor_set(x_53, 3, x_48); +lean_ctor_set(x_53, 4, x_52); +lean_ctor_set(x_53, 5, x_49); +lean_ctor_set(x_53, 6, x_50); +x_54 = lean_st_ref_set(x_7, x_53, x_25); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = lean_box(0); +x_57 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_14, x_16, x_10, x_4, x_5, x_2, x_56, x_6, x_7, x_55); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_90 = x_88; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +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_88); - x_90 = lean_box(0); + lean_dec_ref(x_57); + x_59 = lean_box(0); } -x_91 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; -if (lean_is_scalar(x_90)) { - x_92 = lean_alloc_ctor(0, 2, 0); +x_60 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; +if (lean_is_scalar(x_59)) { + x_61 = lean_alloc_ctor(0, 2, 0); } else { - x_92 = x_90; + x_61 = x_59; } -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_89); -return x_92; +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_58); +return x_61; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_93 = lean_ctor_get(x_88, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_88, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_95 = x_88; +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_57, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_57, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_64 = x_57; } else { - lean_dec_ref(x_88); - x_95 = lean_box(0); + lean_dec_ref(x_57); + x_64 = lean_box(0); } -if (lean_is_scalar(x_95)) { - x_96 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(1, 2, 0); } else { - x_96 = x_95; + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; } -lean_ctor_set(x_96, 0, x_93); -lean_ctor_set(x_96, 1, x_94); -return x_96; } } else { -lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_83, 1); -lean_inc(x_97); -lean_dec(x_83); -lean_inc(x_7); -lean_inc(x_82); -lean_inc(x_2); -x_98 = l_Lean_Compiler_LCNF_ppDecl_x27(x_2, x_82, x_7, x_97); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_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_2); -x_101 = l_Lean_Compiler_LCNF_Decl_size(x_2); -x_102 = l___private_Init_Data_Repr_0__Nat_reprFast(x_101); -x_103 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_103, 0, x_102); -x_104 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_104, 0, x_103); -x_105 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__9; -x_106 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_104); -x_107 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__11; -x_108 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -x_109 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_109, 0, x_99); -x_110 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -x_111 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_110); -lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean_addTrace___at_Lean_Compiler_LCNF_UnreachableBranches_elimDead_go___spec__2(x_10, x_112, x_4, x_5, x_82, x_7, x_100); -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_116 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__1(x_2, x_114, x_4, x_5, x_82, x_7, x_115); -lean_dec(x_114); -if (lean_obj_tag(x_116) == 0) -{ -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_117 = lean_ctor_get(x_116, 1); -lean_inc(x_117); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_118 = x_116; -} else { - lean_dec_ref(x_116); - x_118 = lean_box(0); -} -x_119 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; -if (lean_is_scalar(x_118)) { - x_120 = lean_alloc_ctor(0, 2, 0); -} else { - x_120 = x_118; -} -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_117); -return x_120; +lean_object* x_66; lean_object* x_67; +x_66 = lean_box(0); +x_67 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_14, x_16, x_10, x_4, x_5, x_2, x_66, x_6, x_7, x_19); +if (lean_obj_tag(x_67) == 0) +{ +uint8_t x_68; +x_68 = !lean_is_exclusive(x_67); +if (x_68 == 0) +{ +lean_object* x_69; lean_object* x_70; +x_69 = lean_ctor_get(x_67, 0); +lean_dec(x_69); +x_70 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; +lean_ctor_set(x_67, 0, x_70); +return x_67; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_121 = lean_ctor_get(x_116, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_116, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_123 = x_116; -} else { - lean_dec_ref(x_116); - x_123 = lean_box(0); -} -if (lean_is_scalar(x_123)) { - x_124 = lean_alloc_ctor(1, 2, 0); -} else { - x_124 = x_123; -} -lean_ctor_set(x_124, 0, x_121); -lean_ctor_set(x_124, 1, x_122); -return x_124; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_67, 1); +lean_inc(x_71); +lean_dec(x_67); +x_72 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__9; +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_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -lean_dec(x_82); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -x_125 = lean_ctor_get(x_98, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_98, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_127 = x_98; -} else { - lean_dec_ref(x_98); - x_127 = lean_box(0); +uint8_t x_74; +x_74 = !lean_is_exclusive(x_67); +if (x_74 == 0) +{ +return x_67; } -if (lean_is_scalar(x_127)) { - x_128 = lean_alloc_ctor(1, 2, 0); -} else { - x_128 = x_127; +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_75 = lean_ctor_get(x_67, 0); +x_76 = lean_ctor_get(x_67, 1); +lean_inc(x_76); +lean_inc(x_75); +lean_dec(x_67); +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_ctor_set(x_128, 0, x_125); -lean_ctor_set(x_128, 1, x_126); -return x_128; } } } @@ -1892,7 +2012,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_check _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__1; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -1956,7 +2076,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_1); -x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_1, x_13, x_19, x_6, x_7, x_8, x_9, x_18); +x_20 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3(x_1, x_13, x_19, x_6, x_7, x_8, x_9, x_18); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; @@ -2051,7 +2171,7 @@ lean_dec(x_15); x_38 = lean_ctor_get(x_13, 0); lean_inc(x_38); x_39 = l_Lean_MessageData_ofName(x_38); -x_40 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__13; +x_40 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__7; x_41 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_39); @@ -2083,7 +2203,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_1); -x_53 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_1, x_13, x_51, x_6, x_7, x_8, x_9, x_52); +x_53 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3(x_1, x_13, x_51, x_6, x_7, x_8, x_9, x_52); lean_dec(x_51); if (lean_obj_tag(x_53) == 0) { @@ -2290,11 +2410,21 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_3); return x_9; } @@ -2490,119 +2620,7 @@ x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_4; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; double x_18; double x_19; double x_20; double x_21; double x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_15 = lean_ctor_get(x_13, 0); -x_16 = 0; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Float_ofScientific(x_8, x_16, x_17); -lean_dec(x_8); -x_19 = l_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; -x_20 = lean_float_div(x_18, x_19); -x_21 = l_Float_ofScientific(x_15, x_16, x_17); -lean_dec(x_15); -x_22 = lean_float_div(x_21, x_19); -x_23 = lean_box_float(x_20); -x_24 = lean_box_float(x_22); -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, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_13, 0, x_26); -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_dec(x_13); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_8, x_29, x_30); -lean_dec(x_8); -x_32 = l_Lean_withStartStopSeconds___at_Lean_Compiler_LCNF_PassManager_run___spec__4___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = l_Float_ofScientific(x_27, x_29, x_30); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_32); -x_36 = lean_box_float(x_33); -x_37 = lean_box_float(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_11); -lean_ctor_set(x_39, 1, 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_28); -return x_40; -} -} -else -{ -uint8_t x_41; -lean_dec(x_8); -x_41 = !lean_is_exclusive(x_10); -if (x_41 == 0) -{ -return x_10; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_10, 0); -x_43 = lean_ctor_get(x_10, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_10); -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_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, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; @@ -2747,7 +2765,7 @@ return x_60; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___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) @@ -2772,86 +2790,20 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_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_Compiler_LCNF_PassManager_run___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_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_dec(x_6); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__4(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__6(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_Compiler_LCNF_PassManager_run___spec__5(x_4, x_7, x_8, x_9, x_10, x_13); return x_14; } } -static double _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -2861,7 +2813,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -2869,11 +2821,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___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, uint8_t x_8, double x_9, double 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: { double x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1; +x_17 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_18 = lean_alloc_ctor(0, 2, 17); @@ -2882,7 +2834,7 @@ lean_ctor_set(x_18, 1, x_3); lean_ctor_set_float(x_18, sizeof(void*)*2, x_17); lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_17); lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_2); -x_19 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__2; +x_19 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2___closed__2; x_20 = l_Lean_Option_get___at_Lean_Compiler_LCNF_toConfigOptions___spec__2(x_7, x_19); if (x_20 == 0) { @@ -2892,7 +2844,7 @@ lean_object* x_21; lean_object* x_22; lean_dec(x_3); lean_dec(x_1); x_21 = lean_box(0); -x_22 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); return x_22; } @@ -2907,7 +2859,7 @@ lean_ctor_set_float(x_23, sizeof(void*)*2, x_9); lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_2); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); +x_25 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); return x_25; } @@ -2923,13 +2875,13 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); +x_28 = l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); lean_dec(x_6); return x_28; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Compiler_LCNF_PassManager_run___spec__2___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -2937,16 +2889,16 @@ x_1 = lean_mk_string_from_bytes("", 7); -return x_1; +lean_object* x_6; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_29 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__3; -x_30 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Core_getMaxHeartbeats(x_3); -x_33 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4; -x_34 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_3, x_33); -x_35 = l_Lean_Core_State_nextMacroScope___default___closed__1; -x_36 = l_Lean_Core_State_ngen___default___closed__3; -x_37 = l_Lean_Core_State_traceState___default___closed__3; -x_38 = l_Lean_Core_instInhabitedCache___closed__1; -x_39 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5; -x_40 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_40, 0, x_2); -lean_ctor_set(x_40, 1, x_35); -lean_ctor_set(x_40, 2, x_36); -lean_ctor_set(x_40, 3, x_37); -lean_ctor_set(x_40, 4, x_38); -lean_ctor_set(x_40, 5, x_37); -lean_ctor_set(x_40, 6, x_39); -x_41 = lean_io_get_num_heartbeats(x_6); -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_Core_instMetaEvalCoreM___rarg___lambda__1___closed__6; -x_45 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__2; -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_box(0); -x_48 = lean_box(0); -x_49 = l_Lean_firstFrontendMacroScope; -x_50 = 0; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_inc(x_4); lean_inc(x_3); -x_51 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_51, 0, x_44); -lean_ctor_set(x_51, 1, x_45); -lean_ctor_set(x_51, 2, x_3); -lean_ctor_set(x_51, 3, x_46); -lean_ctor_set(x_51, 4, x_30); -lean_ctor_set(x_51, 5, x_47); -lean_ctor_set(x_51, 6, x_48); -lean_ctor_set(x_51, 7, x_31); -lean_ctor_set(x_51, 8, x_42); -lean_ctor_set(x_51, 9, x_32); -lean_ctor_set(x_51, 10, x_49); -lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_50); -lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_34); -x_52 = lean_st_mk_ref(x_40, x_43); -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_53); -lean_inc(x_51); -x_55 = lean_apply_3(x_4, x_51, x_53, 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_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -lean_inc(x_53); -x_58 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_51, x_53, x_57); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = lean_st_ref_get(x_53, x_59); -lean_dec(x_53); -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_ctor_get(x_61, 0); -lean_inc(x_63); -lean_dec(x_61); -x_64 = 1; -x_65 = lean_box(x_64); -x_66 = lean_apply_5(x_1, x_63, x_3, x_56, x_65, x_62); -return x_66; -} -else +lean_inc(x_2); +x_7 = l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3(x_2, x_6, x_2, x_3, x_4, x_5); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_67; lean_object* x_68; -lean_dec(x_56); -lean_dec(x_53); -lean_dec(x_3); -lean_dec(x_1); -x_67 = lean_ctor_get(x_58, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_58, 1); -lean_inc(x_68); -lean_dec(x_58); -x_7 = x_67; -x_8 = x_68; -goto block_28; -} -} -else +lean_object* x_8; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_69; lean_object* x_70; lean_object* x_71; +uint8_t x_9; +lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_69 = lean_ctor_get(x_55, 0); -lean_inc(x_69); -x_70 = lean_ctor_get(x_55, 1); -lean_inc(x_70); -lean_dec(x_55); -x_71 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_51, x_53, x_70); -if (lean_obj_tag(x_71) == 0) +x_9 = !lean_is_exclusive(x_7); +if (x_9 == 0) { -lean_object* x_72; -x_72 = lean_ctor_get(x_71, 1); -lean_inc(x_72); -lean_dec(x_71); -x_7 = x_69; -x_8 = x_72; -goto block_28; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_7, 0); +lean_dec(x_10); +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +lean_dec(x_8); +lean_ctor_set(x_7, 0, x_11); +return x_7; } else { -lean_object* x_73; lean_object* x_74; -lean_dec(x_69); -x_73 = lean_ctor_get(x_71, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_71, 1); -lean_inc(x_74); -lean_dec(x_71); -x_7 = x_73; -x_8 = x_74; -goto block_28; +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_dec(x_7); +x_13 = lean_ctor_get(x_8, 0); +lean_inc(x_13); +lean_dec(x_8); +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; } } -block_28: -{ -if (lean_obj_tag(x_7) == 0) +else { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); +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_15 = lean_ctor_get(x_7, 1); +lean_inc(x_15); lean_dec(x_7); -x_10 = l_Lean_MessageData_toString(x_9, x_8); -if (lean_obj_tag(x_10) == 0) +x_16 = lean_ctor_get(x_8, 0); +lean_inc(x_16); +lean_dec(x_8); +x_17 = lean_ctor_get(x_1, 1); +lean_inc(x_17); +lean_dec(x_1); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_16); +x_20 = lean_array_get_size(x_17); +x_21 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_22 = 0; +x_23 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__6(x_18, x_17, x_21, x_22, x_19, x_3, x_4, x_15); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_17); +if (lean_obj_tag(x_23) == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set_tag(x_10, 1); -lean_ctor_set(x_10, 0, x_13); -return x_10; -} -else +uint8_t x_26; +x_26 = !lean_is_exclusive(x_23); +if (x_26 == 0) { -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); -x_16 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_16, 0, 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_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_dec(x_27); +x_28 = lean_ctor_get(x_24, 1); +lean_inc(x_28); +lean_dec(x_24); +lean_ctor_set(x_23, 0, x_28); +return x_23; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_10); -if (x_18 == 0) -{ -return x_10; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +lean_dec(x_23); +x_30 = lean_ctor_get(x_24, 1); +lean_inc(x_30); +lean_dec(x_24); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; +} } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_10, 0); -x_20 = lean_ctor_get(x_10, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_10); -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; -} -} +uint8_t x_32; +lean_dec(x_24); +x_32 = !lean_is_exclusive(x_23); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_23, 0); +lean_dec(x_33); +x_34 = lean_ctor_get(x_25, 0); +lean_inc(x_34); +lean_dec(x_25); +lean_ctor_set(x_23, 0, x_34); +return x_23; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_7, 0); -lean_inc(x_22); -lean_dec(x_7); -x_23 = l___private_Init_Data_Repr_0__Nat_reprFast(x_22); -x_24 = l_Lean_Core_CoreM_toIO___rarg___closed__1; -x_25 = lean_string_append(x_24, x_23); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_23, 1); +lean_inc(x_35); lean_dec(x_23); -x_26 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_26, 0, x_25); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_8); -return x_27; +x_36 = lean_ctor_get(x_25, 0); +lean_inc(x_36); +lean_dec(x_25); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +return x_37; } } } +else +{ +uint8_t x_38; +x_38 = !lean_is_exclusive(x_23); +if (x_38 == 0) +{ +return x_23; } -LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___boxed), 6, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_23, 0); +x_40 = lean_ctor_get(x_23, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_23); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } -LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Core_instMetaEvalCoreM___rarg), 1, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___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: +else { -size_t x_10; size_t x_11; lean_object* x_12; -x_10 = lean_unbox_usize(x_4); +uint8_t x_42; lean_dec(x_4); -x_11 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_12 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__4(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_dec(x_1); +x_42 = !lean_is_exclusive(x_7); +if (x_42 == 0) +{ +return x_7; } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___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: +else { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_10 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_11 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__5(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_7, 0); +x_44 = lean_ctor_get(x_7, 1); +lean_inc(x_44); +lean_inc(x_43); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -return 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_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___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) { -_start: -{ -lean_object* x_6; -x_6 = l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3___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_Core_instMetaEvalCoreM___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_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_9; size_t x_10; lean_object* x_11; -x_9 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_10 = lean_unbox_usize(x_4); +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_4 = lean_st_ref_get(x_2, x_3); +x_5 = lean_ctor_get(x_4, 0); +lean_inc(x_5); +x_6 = lean_ctor_get(x_4, 1); +lean_inc(x_6); lean_dec(x_4); -x_11 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__6(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_2); -return x_11; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___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) { -_start: +x_7 = lean_ctor_get(x_5, 3); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_box(0); +x_9 = l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2(x_7, x_8, x_1, x_2, x_6); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_6; -x_6 = l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2___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; -} +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_9, 0); +lean_dec(x_11); +lean_ctor_set(x_9, 0, x_8); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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: +else { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_5); -lean_dec(x_5); -x_8 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_2, x_3, x_4, x_7, x_6); -return x_8; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_8); +lean_ctor_set(x_13, 1, x_12); +return x_13; } } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_maxRecDepthErrorMessage; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +uint8_t x_14; +x_14 = !lean_is_exclusive(x_9); +if (x_14 == 0) +{ +return x_9; } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_9, 0); +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_9); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_1); -lean_ctor_set(x_6, 1, x_5); -x_7 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_7, 0, x_6); -lean_ctor_set(x_7, 1, x_4); -return x_7; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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_object* x_7, lean_object* x_8) { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed), 4, 0); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___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) { -_start: +uint8_t x_9; +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; -lean_dec(x_2); -x_7 = lean_ctor_get(x_4, 0); +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_6, 4); +lean_dec(x_10); +x_11 = lean_ctor_get(x_6, 2); +lean_dec(x_11); +lean_ctor_set(x_6, 4, x_2); +lean_ctor_set(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, sizeof(void*)*11 + 1, x_3); lean_inc(x_7); -x_8 = lean_ctor_get(x_4, 1); -lean_inc(x_8); -x_9 = lean_ctor_get(x_4, 2); -lean_inc(x_9); -x_10 = lean_ctor_get(x_4, 3); -lean_inc(x_10); -x_11 = lean_ctor_get(x_4, 4); -lean_inc(x_11); -x_12 = lean_ctor_get(x_4, 5); -lean_inc(x_12); -x_13 = lean_ctor_get(x_4, 6); +lean_inc(x_6); +x_12 = lean_apply_3(x_4, 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; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -x_14 = lean_ctor_get(x_4, 7); +x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); -x_15 = lean_ctor_get(x_4, 8); -lean_inc(x_15); -x_16 = lean_ctor_get(x_4, 9); -lean_inc(x_16); -x_17 = lean_ctor_get(x_4, 10); -lean_inc(x_17); -x_18 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); -x_19 = lean_ctor_get_uint8(x_4, sizeof(void*)*11 + 1); -x_20 = lean_nat_dec_eq(x_10, x_11); -if (x_20 == 0) +lean_dec(x_12); +x_15 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_6, x_7, x_14); +if (lean_obj_tag(x_15) == 0) { -uint8_t x_21; -x_21 = !lean_is_exclusive(x_4); -if (x_21 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 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; -x_22 = lean_ctor_get(x_4, 10); -lean_dec(x_22); -x_23 = lean_ctor_get(x_4, 9); -lean_dec(x_23); -x_24 = lean_ctor_get(x_4, 8); -lean_dec(x_24); -x_25 = lean_ctor_get(x_4, 7); -lean_dec(x_25); -x_26 = lean_ctor_get(x_4, 6); -lean_dec(x_26); -x_27 = lean_ctor_get(x_4, 5); -lean_dec(x_27); -x_28 = lean_ctor_get(x_4, 4); -lean_dec(x_28); -x_29 = lean_ctor_get(x_4, 3); -lean_dec(x_29); -x_30 = lean_ctor_get(x_4, 2); -lean_dec(x_30); -x_31 = lean_ctor_get(x_4, 1); -lean_dec(x_31); -x_32 = lean_ctor_get(x_4, 0); -lean_dec(x_32); -x_33 = lean_unsigned_to_nat(1u); -x_34 = lean_nat_add(x_10, x_33); -lean_dec(x_10); -lean_ctor_set(x_4, 3, x_34); -x_35 = lean_apply_5(x_3, lean_box(0), x_1, x_4, x_5, x_6); -return x_35; +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +lean_ctor_set(x_15, 0, x_13); +return x_15; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_4); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_10, x_36); -lean_dec(x_10); -x_38 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_38, 0, x_7); -lean_ctor_set(x_38, 1, x_8); -lean_ctor_set(x_38, 2, x_9); -lean_ctor_set(x_38, 3, x_37); -lean_ctor_set(x_38, 4, x_11); -lean_ctor_set(x_38, 5, x_12); -lean_ctor_set(x_38, 6, x_13); -lean_ctor_set(x_38, 7, x_14); -lean_ctor_set(x_38, 8, x_15); -lean_ctor_set(x_38, 9, x_16); -lean_ctor_set(x_38, 10, x_17); -lean_ctor_set_uint8(x_38, sizeof(void*)*11, x_18); -lean_ctor_set_uint8(x_38, sizeof(void*)*11 + 1, x_19); -x_39 = lean_apply_5(x_3, lean_box(0), x_1, x_38, x_5, x_6); -return x_39; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } else { -lean_object* x_40; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_20; 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_3); -lean_dec(x_1); -x_40 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg(x_12, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -return x_40; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +x_20 = !lean_is_exclusive(x_15); +if (x_20 == 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; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -lean_inc(x_2); -x_5 = lean_alloc_closure((void*)(l_Lean_Core_withIncRecDepth___rarg___lambda__1), 6, 2); -lean_closure_set(x_5, 0, x_3); -lean_closure_set(x_5, 1, x_2); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_apply_2(x_6, lean_box(0), x_5); -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_apply_1(x_8, lean_box(0)); -x_10 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_9); -return x_10; -} +return x_15; } -LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Core_withIncRecDepth___rarg), 3, 0); -return x_3; -} +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_15, 0); +x_22 = lean_ctor_get(x_15, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_15); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; -} +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_12, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_12, 1); +lean_inc(x_25); +lean_dec(x_12); +x_26 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_6, x_7, x_25); +if (lean_obj_tag(x_26) == 0) +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +lean_object* x_28; +x_28 = lean_ctor_get(x_26, 0); +lean_dec(x_28); +lean_ctor_set_tag(x_26, 1); +lean_ctor_set(x_26, 0, x_24); +return x_26; } -static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("elaboration interrupted", 23); -return x_1; +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(1, 2, 0); +lean_ctor_set(x_30, 0, x_24); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Core_checkInterrupted___rarg___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} +uint8_t x_31; +lean_dec(x_24); +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) +{ +return x_26; } -static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Core_checkInterrupted___rarg___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_26); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } -static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Core_checkInterrupted___rarg___closed__3; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = lean_io_check_canceled(x_1); -x_3 = lean_ctor_get(x_2, 0); -lean_inc(x_3); -x_4 = lean_unbox(x_3); -lean_dec(x_3); -if (x_4 == 0) +else { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_2); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_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; uint8_t x_44; lean_object* x_45; lean_object* x_46; +x_35 = lean_ctor_get(x_6, 0); +x_36 = lean_ctor_get(x_6, 1); +x_37 = lean_ctor_get(x_6, 3); +x_38 = lean_ctor_get(x_6, 5); +x_39 = lean_ctor_get(x_6, 6); +x_40 = lean_ctor_get(x_6, 7); +x_41 = lean_ctor_get(x_6, 8); +x_42 = lean_ctor_get(x_6, 9); +x_43 = lean_ctor_get(x_6, 10); +x_44 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); lean_dec(x_6); -x_7 = lean_box(0); -lean_ctor_set(x_2, 0, x_7); -return x_2; -} -else +x_45 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_45, 0, x_35); +lean_ctor_set(x_45, 1, x_36); +lean_ctor_set(x_45, 2, x_1); +lean_ctor_set(x_45, 3, x_37); +lean_ctor_set(x_45, 4, x_2); +lean_ctor_set(x_45, 5, x_38); +lean_ctor_set(x_45, 6, x_39); +lean_ctor_set(x_45, 7, x_40); +lean_ctor_set(x_45, 8, x_41); +lean_ctor_set(x_45, 9, x_42); +lean_ctor_set(x_45, 10, x_43); +lean_ctor_set_uint8(x_45, sizeof(void*)*11, x_44); +lean_ctor_set_uint8(x_45, sizeof(void*)*11 + 1, x_3); +lean_inc(x_7); +lean_inc(x_45); +x_46 = lean_apply_3(x_4, x_45, x_7, x_8); +if (lean_obj_tag(x_46) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +x_49 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_45, x_7, x_48); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_49, 1); +lean_inc(x_50); +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); } +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(0, 2, 0); +} else { + x_52 = x_51; } -else -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_2, 0); -lean_dec(x_12); -x_13 = l_Lean_Core_checkInterrupted___rarg___closed__4; -lean_ctor_set_tag(x_2, 1); -lean_ctor_set(x_2, 0, x_13); -return x_2; +lean_ctor_set(x_52, 0, x_47); +lean_ctor_set(x_52, 1, x_50); +return x_52; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_2, 1); -lean_inc(x_14); -lean_dec(x_2); -x_15 = l_Lean_Core_checkInterrupted___rarg___closed__4; -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +lean_dec(x_47); +x_53 = lean_ctor_get(x_49, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_49, 1); +lean_inc(x_54); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_55 = x_49; +} else { + lean_dec_ref(x_49); + x_55 = lean_box(0); } +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); +} else { + x_56 = x_55; } +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; } } -LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Core_checkInterrupted___rarg), 1, 0); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted___boxed(lean_object* x_1, lean_object* x_2) { -_start: +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_46, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_46, 1); +lean_inc(x_58); +lean_dec(x_46); +x_59 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_45, x_7, x_58); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_3; -x_3 = l_Lean_Core_checkInterrupted(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -return x_3; -} +lean_object* x_60; lean_object* x_61; lean_object* x_62; +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; +} else { + lean_dec_ref(x_59); + x_61 = lean_box(0); } -static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(deterministic) timeout at '", 28); -return x_1; +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(1, 2, 0); +} else { + x_62 = x_61; + lean_ctor_set_tag(x_62, 1); } +lean_ctor_set(x_62, 0, x_57); +lean_ctor_set(x_62, 1, x_60); +return x_62; } -static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("', maximum number of heartbeats (", 33); -return x_1; -} +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_dec(x_57); +x_63 = lean_ctor_get(x_59, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_59, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_65 = x_59; +} else { + lean_dec_ref(x_59); + x_65 = lean_box(0); } -static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(") has been reached (use 'set_option ", 36); -return x_1; +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(1, 2, 0); +} else { + x_66 = x_65; } +lean_ctor_set(x_66, 0, x_63); +lean_ctor_set(x_66, 1, x_64); +return x_66; } -static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" ' to set the limit)", 25); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat(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; lean_object* x_9; lean_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; -x_7 = 1; -x_8 = l_Lean_Name_toString(x_1, x_7); -x_9 = l_Lean_Core_throwMaxHeartbeat___closed__1; -x_10 = lean_string_append(x_9, x_8); -lean_dec(x_8); -x_11 = l_Lean_Core_throwMaxHeartbeat___closed__2; -x_12 = lean_string_append(x_10, x_11); -x_13 = lean_unsigned_to_nat(1000u); -x_14 = lean_nat_div(x_3, x_13); -x_15 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); -x_16 = lean_string_append(x_12, x_15); -lean_dec(x_15); -x_17 = l_Lean_Core_throwMaxHeartbeat___closed__3; -x_18 = lean_string_append(x_16, x_17); -x_19 = l_Lean_Name_toString(x_2, x_7); -x_20 = lean_string_append(x_18, x_19); -lean_dec(x_19); -x_21 = l_Lean_Core_throwMaxHeartbeat___closed__4; -x_22 = lean_string_append(x_20, x_21); -x_23 = lean_ctor_get(x_4, 5); -x_24 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_24, 0, x_22); -x_25 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_25, 0, x_24); -lean_inc(x_23); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_23); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_6); -return x_27; } } -LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat___boxed(lean_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_Core_instMetaEvalCoreM___rarg___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_7; -x_7 = l_Lean_Core_throwMaxHeartbeat(x_1, x_2, x_3, x_4, x_5, x_6); +uint8_t x_9; lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_unsigned_to_nat(0u); -x_8 = lean_nat_dec_eq(x_3, x_7); -if (x_8 == 0) -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_io_get_num_heartbeats(x_6); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_4, 8); -x_14 = lean_nat_sub(x_11, 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; lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; +x_10 = lean_ctor_get(x_6, 4); +lean_dec(x_10); +x_11 = lean_ctor_get(x_6, 2); lean_dec(x_11); -x_15 = lean_nat_dec_lt(x_3, x_14); +x_12 = l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1; +x_13 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_12); +lean_inc(x_13); +lean_inc(x_1); +lean_ctor_set(x_6, 4, x_13); +lean_ctor_set(x_6, 2, x_1); +lean_ctor_set_uint8(x_6, sizeof(void*)*11 + 1, x_2); +x_14 = lean_st_ref_get(x_7, x_8); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); lean_dec(x_14); -if (x_15 == 0) +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Kernel_isDiagnosticsEnabled(x_17); +lean_dec(x_17); +if (x_18 == 0) { -lean_object* x_16; -lean_dec(x_2); -lean_dec(x_1); -x_16 = lean_box(0); -lean_ctor_set(x_9, 0, x_16); -return x_9; +if (x_2 == 0) +{ +uint8_t x_46; +x_46 = 1; +x_19 = x_46; +goto block_45; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -lean_free_object(x_9); -x_17 = lean_box(0); -x_18 = l_Lean_Name_str___override(x_17, x_1); -x_19 = l_Lean_Core_throwMaxHeartbeat(x_18, x_2, x_3, x_4, x_5, x_12); -return x_19; +uint8_t x_47; +x_47 = 0; +x_19 = x_47; +goto block_45; } } 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_9, 0); -x_21 = lean_ctor_get(x_9, 1); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_9); -x_22 = lean_ctor_get(x_4, 8); -x_23 = lean_nat_sub(x_20, x_22); -lean_dec(x_20); -x_24 = lean_nat_dec_lt(x_3, x_23); -lean_dec(x_23); -if (x_24 == 0) +if (x_2 == 0) { -lean_object* x_25; lean_object* x_26; -lean_dec(x_2); -lean_dec(x_1); -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_21); -return x_26; +uint8_t x_48; +x_48 = 0; +x_19 = x_48; +goto block_45; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_box(0); -x_28 = l_Lean_Name_str___override(x_27, x_1); -x_29 = l_Lean_Core_throwMaxHeartbeat(x_28, x_2, x_3, x_4, x_5, x_21); -return x_29; -} +uint8_t x_49; +x_49 = 1; +x_19 = x_49; +goto block_45; } } -else +block_45: { -lean_object* x_30; lean_object* x_31; -lean_dec(x_2); -lean_dec(x_1); -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_6); -return x_31; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore___boxed(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 (x_19 == 0) { -lean_object* x_7; -x_7 = l_Lean_Core_checkMaxHeartbeatsCore(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_7; -} +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_st_ref_take(x_7, x_16); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_21, 0); +x_25 = lean_ctor_get(x_21, 4); +lean_dec(x_25); +x_26 = l_Lean_Kernel_enableDiag(x_24, x_2); +lean_ctor_set(x_21, 4, x_4); +lean_ctor_set(x_21, 0, x_26); +x_27 = lean_st_ref_set(x_7, x_21, x_22); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_13, x_2, x_3, x_29, x_6, x_7, x_28); +return x_30; } -LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_2, 9); -x_6 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__2; -x_7 = l_Lean_Core_checkMaxHeartbeatsCore(x_1, x_6, x_5, x_2, x_3, x_4); -return x_7; +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_31 = lean_ctor_get(x_21, 0); +x_32 = lean_ctor_get(x_21, 1); +x_33 = lean_ctor_get(x_21, 2); +x_34 = lean_ctor_get(x_21, 3); +x_35 = lean_ctor_get(x_21, 5); +x_36 = lean_ctor_get(x_21, 6); +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_21); +x_37 = l_Lean_Kernel_enableDiag(x_31, x_2); +x_38 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_32); +lean_ctor_set(x_38, 2, x_33); +lean_ctor_set(x_38, 3, x_34); +lean_ctor_set(x_38, 4, x_4); +lean_ctor_set(x_38, 5, x_35); +lean_ctor_set(x_38, 6, x_36); +x_39 = lean_st_ref_set(x_7, x_38, x_22); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_box(0); +x_42 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_13, x_2, x_3, x_41, x_6, x_7, x_40); +return x_42; } } -LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats___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_Core_checkMaxHeartbeats(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +lean_object* x_43; lean_object* x_44; +lean_dec(x_4); +x_43 = lean_box(0); +x_44 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_13, x_2, x_3, x_43, x_6, x_7, x_16); +return x_44; } } -LEAN_EXPORT lean_object* l_Lean_Core_checkSystem(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_io_check_canceled(x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_unbox(x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); -lean_dec(x_5); -x_9 = l_Lean_Core_checkMaxHeartbeats(x_1, x_2, x_3, x_8); -return x_9; } else { -uint8_t x_10; -lean_dec(x_1); -x_10 = !lean_is_exclusive(x_5); -if (x_10 == 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; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; uint8_t x_68; +x_50 = lean_ctor_get(x_6, 0); +x_51 = lean_ctor_get(x_6, 1); +x_52 = lean_ctor_get(x_6, 3); +x_53 = lean_ctor_get(x_6, 5); +x_54 = lean_ctor_get(x_6, 6); +x_55 = lean_ctor_get(x_6, 7); +x_56 = lean_ctor_get(x_6, 8); +x_57 = lean_ctor_get(x_6, 9); +x_58 = lean_ctor_get(x_6, 10); +x_59 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +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_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_6); +x_60 = l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1; +x_61 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_60); +lean_inc(x_61); +lean_inc(x_1); +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_50); +lean_ctor_set(x_62, 1, x_51); +lean_ctor_set(x_62, 2, x_1); +lean_ctor_set(x_62, 3, x_52); +lean_ctor_set(x_62, 4, x_61); +lean_ctor_set(x_62, 5, x_53); +lean_ctor_set(x_62, 6, x_54); +lean_ctor_set(x_62, 7, x_55); +lean_ctor_set(x_62, 8, x_56); +lean_ctor_set(x_62, 9, x_57); +lean_ctor_set(x_62, 10, x_58); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_59); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_2); +x_63 = lean_st_ref_get(x_7, x_8); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_ctor_get(x_64, 0); +lean_inc(x_66); +lean_dec(x_64); +x_67 = l_Lean_Kernel_isDiagnosticsEnabled(x_66); +lean_dec(x_66); +if (x_67 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_5, 0); -lean_dec(x_11); -x_12 = l_Lean_Core_checkInterrupted___rarg___closed__4; -lean_ctor_set_tag(x_5, 1); -lean_ctor_set(x_5, 0, x_12); -return x_5; +if (x_2 == 0) +{ +uint8_t x_88; +x_88 = 1; +x_68 = x_88; +goto block_87; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_5, 1); -lean_inc(x_13); -lean_dec(x_5); -x_14 = l_Lean_Core_checkInterrupted___rarg___closed__4; -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; -} +uint8_t x_89; +x_89 = 0; +x_68 = x_89; +goto block_87; } } +else +{ +if (x_2 == 0) +{ +uint8_t x_90; +x_90 = 0; +x_68 = x_90; +goto block_87; } -LEAN_EXPORT lean_object* l_Lean_Core_checkSystem___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_Core_checkSystem(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +uint8_t x_91; +x_91 = 1; +x_68 = x_91; +goto block_87; } } -LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +block_87: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_io_get_num_heartbeats(x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = !lean_is_exclusive(x_2); -if (x_8 == 0) +if (x_68 == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_2, 8); -lean_dec(x_9); -lean_ctor_set(x_2, 8, x_6); -x_10 = lean_apply_3(x_1, x_2, x_3, x_7); -return x_10; +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_69 = lean_st_ref_take(x_7, x_65); +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, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_70, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_70, 2); +lean_inc(x_74); +x_75 = lean_ctor_get(x_70, 3); +lean_inc(x_75); +x_76 = lean_ctor_get(x_70, 5); +lean_inc(x_76); +x_77 = lean_ctor_get(x_70, 6); +lean_inc(x_77); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + lean_ctor_release(x_70, 2); + lean_ctor_release(x_70, 3); + lean_ctor_release(x_70, 4); + lean_ctor_release(x_70, 5); + lean_ctor_release(x_70, 6); + x_78 = x_70; +} else { + lean_dec_ref(x_70); + x_78 = lean_box(0); +} +x_79 = l_Lean_Kernel_enableDiag(x_72, x_2); +if (lean_is_scalar(x_78)) { + x_80 = lean_alloc_ctor(0, 7, 0); +} else { + x_80 = x_78; +} +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_73); +lean_ctor_set(x_80, 2, x_74); +lean_ctor_set(x_80, 3, x_75); +lean_ctor_set(x_80, 4, x_4); +lean_ctor_set(x_80, 5, x_76); +lean_ctor_set(x_80, 6, x_77); +x_81 = lean_st_ref_set(x_7, x_80, x_71); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +x_83 = lean_box(0); +x_84 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_61, x_2, x_3, x_83, x_62, x_7, x_82); +return x_84; } 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; uint8_t x_22; lean_object* x_23; lean_object* x_24; -x_11 = lean_ctor_get(x_2, 0); -x_12 = lean_ctor_get(x_2, 1); -x_13 = lean_ctor_get(x_2, 2); -x_14 = lean_ctor_get(x_2, 3); -x_15 = lean_ctor_get(x_2, 4); -x_16 = lean_ctor_get(x_2, 5); -x_17 = lean_ctor_get(x_2, 6); -x_18 = lean_ctor_get(x_2, 7); -x_19 = lean_ctor_get(x_2, 9); -x_20 = lean_ctor_get(x_2, 10); -x_21 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); -x_22 = lean_ctor_get_uint8(x_2, sizeof(void*)*11 + 1); -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_12); -lean_inc(x_11); -lean_dec(x_2); -x_23 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_23, 0, x_11); -lean_ctor_set(x_23, 1, x_12); -lean_ctor_set(x_23, 2, x_13); -lean_ctor_set(x_23, 3, x_14); -lean_ctor_set(x_23, 4, x_15); -lean_ctor_set(x_23, 5, x_16); -lean_ctor_set(x_23, 6, x_17); -lean_ctor_set(x_23, 7, x_18); -lean_ctor_set(x_23, 8, x_6); -lean_ctor_set(x_23, 9, x_19); -lean_ctor_set(x_23, 10, x_20); -lean_ctor_set_uint8(x_23, sizeof(void*)*11, x_21); -lean_ctor_set_uint8(x_23, sizeof(void*)*11 + 1, x_22); -x_24 = lean_apply_3(x_1, x_23, x_3, x_7); -return x_24; +lean_object* x_85; lean_object* x_86; +lean_dec(x_4); +x_85 = lean_box(0); +x_86 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_61, x_2, x_3, x_85, x_62, x_7, x_65); +return x_86; } } } -LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp(lean_object* x_1) { +} +} +static lean_object* _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg), 4, 0); +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg___lambda__1(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_Core_instMetaEvalCoreM___rarg___closed__2() { _start: { -lean_object* x_6; lean_object* x_7; -x_6 = lean_apply_2(x_2, lean_box(0), x_1); -x_7 = l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(x_6, x_3, x_4, x_5); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; +x_2 = l_Lean_Core_instMetaEvalCoreM___rarg___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__3() { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -lean_dec(x_1); -x_5 = lean_alloc_closure((void*)(l_Lean_Core_withCurrHeartbeats___rarg___lambda__1), 5, 1); -lean_closure_set(x_5, 0, x_3); -x_6 = lean_ctor_get(x_2, 0); -lean_inc(x_6); -x_7 = lean_apply_2(x_6, lean_box(0), x_5); -x_8 = lean_ctor_get(x_2, 1); -lean_inc(x_8); -lean_dec(x_2); -x_9 = lean_apply_1(x_8, lean_box(0)); -x_10 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_9); -return x_10; +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_Core_State_infoState___default___closed__1; +x_3 = l_Lean_Core_State_traceState___default___closed__3; +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; } } -LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__4() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Core_withCurrHeartbeats___rarg), 3, 0); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 7); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_st_ref_take(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = !lean_is_exclusive(x_6); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_6, 5); -lean_dec(x_9); -lean_ctor_set(x_6, 5, x_1); -x_10 = lean_st_ref_set(x_3, x_6, x_7); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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) { +_start: +{ +lean_object* x_7; lean_object* x_21; lean_object* x_22; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; +x_43 = lean_box(0); +x_44 = l_Lean_Core_getMaxHeartbeats(x_3); +x_45 = l_Lean_Core_State_nextMacroScope___default___closed__1; +x_46 = l_Lean_Core_State_ngen___default___closed__3; +x_47 = l_Lean_Core_State_traceState___default___closed__3; +x_48 = l_Lean_Core_instInhabitedCache___closed__1; +x_49 = l_Lean_Core_instMetaEvalCoreM___rarg___closed__3; +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_45); +lean_ctor_set(x_50, 2, x_46); +lean_ctor_set(x_50, 3, x_47); +lean_ctor_set(x_50, 4, x_48); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_49); +x_51 = lean_io_get_num_heartbeats(x_6); +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_Core_instMetaEvalCoreM___rarg___closed__4; +x_55 = l_Lean_Core_instMetaEvalCoreM___rarg___closed__2; +x_56 = lean_unsigned_to_nat(0u); +x_57 = lean_unsigned_to_nat(1000u); +x_58 = lean_box(0); +x_59 = lean_box(0); +x_60 = l_Lean_firstFrontendMacroScope; +x_61 = 0; +lean_inc(x_3); +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_54); +lean_ctor_set(x_62, 1, x_55); +lean_ctor_set(x_62, 2, x_3); +lean_ctor_set(x_62, 3, x_56); +lean_ctor_set(x_62, 4, x_57); +lean_ctor_set(x_62, 5, x_58); +lean_ctor_set(x_62, 6, x_59); +lean_ctor_set(x_62, 7, x_43); +lean_ctor_set(x_62, 8, x_52); +lean_ctor_set(x_62, 9, x_44); +lean_ctor_set(x_62, 10, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_61); +x_63 = lean_st_mk_ref(x_50, x_53); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1; +x_67 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_3, x_66); +x_68 = lean_st_ref_get(x_64, x_65); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +lean_dec(x_69); +x_72 = l_Lean_Kernel_isDiagnosticsEnabled(x_71); +lean_dec(x_71); +if (x_72 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -lean_dec(x_12); -x_13 = lean_box(0); -lean_ctor_set(x_10, 0, x_13); -return x_10; +if (x_67 == 0) +{ +uint8_t x_134; +x_134 = 1; +x_73 = x_134; +goto block_133; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_10, 1); -lean_inc(x_14); -lean_dec(x_10); -x_15 = lean_box(0); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +x_73 = x_61; +goto block_133; } } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -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, 6); -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_dec(x_6); -x_23 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_23, 0, x_17); -lean_ctor_set(x_23, 1, x_18); -lean_ctor_set(x_23, 2, x_19); -lean_ctor_set(x_23, 3, x_20); -lean_ctor_set(x_23, 4, x_21); -lean_ctor_set(x_23, 5, x_1); -lean_ctor_set(x_23, 6, x_22); -x_24 = lean_st_ref_set(x_3, x_23, x_7); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -if (lean_is_exclusive(x_24)) { - lean_ctor_release(x_24, 0); - lean_ctor_release(x_24, 1); - x_26 = x_24; -} else { - lean_dec_ref(x_24); - x_26 = lean_box(0); -} -x_27 = lean_box(0); -if (lean_is_scalar(x_26)) { - x_28 = lean_alloc_ctor(0, 2, 0); -} else { - x_28 = x_26; -} -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_25); -return x_28; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +if (x_67 == 0) { -lean_object* x_5; -x_5 = l_Lean_Core_setMessageLog(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; +x_73 = x_61; +goto block_133; } -} -LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog(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_Core_State_traceState___default___closed__3; -x_5 = l_Lean_Core_setMessageLog(x_4, x_1, x_2, x_3); -return x_5; +uint8_t x_135; +x_135 = 1; +x_73 = x_135; +goto block_133; } } -LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +block_20: { -lean_object* x_4; -x_4 = l_Lean_Core_resetMessageLog(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; -} +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_8, 1); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = 1; +x_14 = lean_box(x_13); +x_15 = lean_apply_5(x_1, x_12, x_3, x_10, x_14, x_9); +return x_15; } -LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___rarg(lean_object* x_1, lean_object* x_2) { -_start: +else { -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) +uint8_t x_16; +lean_dec(x_3); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_7); +if (x_16 == 0) { -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_3, 0); -x_6 = lean_ctor_get(x_5, 5); -lean_inc(x_6); -lean_dec(x_5); -lean_ctor_set(x_3, 0, x_6); -return x_3; +return x_7; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_3, 0); -x_8 = lean_ctor_get(x_3, 1); -lean_inc(x_8); -lean_inc(x_7); -lean_dec(x_3); -x_9 = lean_ctor_get(x_7, 5); -lean_inc(x_9); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_7, 0); +x_18 = lean_ctor_get(x_7, 1); +lean_inc(x_18); +lean_inc(x_17); lean_dec(x_7); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog(lean_object* x_1) { -_start: +block_42: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Core_getMessageLog___rarg___boxed), 2, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: +if (lean_obj_tag(x_21) == 0) { -lean_object* x_3; -x_3 = l_Lean_Core_getMessageLog___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_MessageData_toString(x_23, x_22); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_tag(x_24, 1); +lean_ctor_set(x_24, 0, x_27); +x_7 = x_24; +goto block_20; } -LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___boxed(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = l_Lean_Core_getMessageLog(x_1); -lean_dec(x_1); -return x_2; +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 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_30, 0, x_28); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_7 = x_31; +goto block_20; } } -LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___rarg(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = lean_st_ref_take(x_1, x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec(x_3); -x_6 = lean_st_ref_set(x_1, x_4, x_5); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_dec(x_8); -x_9 = l_Lean_Core_State_traceState___default___closed__3; -lean_ctor_set(x_6, 0, x_9); -return x_6; +x_7 = x_24; +goto block_20; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_6, 1); -lean_inc(x_10); -lean_dec(x_6); -x_11 = l_Lean_Core_State_traceState___default___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_10); -return x_12; +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); +x_7 = x_35; +goto block_20; } } } -LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Core_getAndEmptyMessageLog___rarg___boxed), 2, 0); -return x_2; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_21, 0); +lean_inc(x_36); +lean_dec(x_21); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_36); +x_38 = l_Lean_Core_CoreM_toIO___rarg___closed__1; +x_39 = lean_string_append(x_38, x_37); +lean_dec(x_37); +x_40 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_22); +x_7 = x_41; +goto block_20; } } -LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: +block_133: { -lean_object* x_3; -x_3 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} +if (x_73 == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = lean_st_ref_take(x_64, x_70); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = !lean_is_exclusive(x_75); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_78 = lean_ctor_get(x_75, 0); +x_79 = lean_ctor_get(x_75, 4); +lean_dec(x_79); +x_80 = l_Lean_Kernel_enableDiag(x_78, x_67); +lean_ctor_set(x_75, 4, x_48); +lean_ctor_set(x_75, 0, x_80); +x_81 = lean_st_ref_set(x_64, x_75, x_76); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +x_83 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_84 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__2(x_3, x_67, x_4, x_48, x_83, x_62, x_64, x_82); +if (lean_obj_tag(x_84) == 0) +{ +lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; +x_85 = lean_ctor_get(x_84, 0); +lean_inc(x_85); +x_86 = lean_ctor_get(x_84, 1); +lean_inc(x_86); +lean_dec(x_84); +x_87 = lean_st_ref_get(x_64, x_86); +lean_dec(x_64); +x_88 = !lean_is_exclusive(x_87); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; +x_89 = lean_ctor_get(x_87, 0); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_85); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_87, 0, x_90); +x_7 = x_87; +goto block_20; } -LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___boxed(lean_object* x_1) { +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_91 = lean_ctor_get(x_87, 0); +x_92 = lean_ctor_get(x_87, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_87); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_85); +lean_ctor_set(x_93, 1, x_91); +x_94 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_92); +x_7 = x_94; +goto block_20; +} +} +else +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_64); +x_95 = lean_ctor_get(x_84, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_84, 1); +lean_inc(x_96); +lean_dec(x_84); +x_21 = x_95; +x_22 = x_96; +goto block_42; +} +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_97 = lean_ctor_get(x_75, 0); +x_98 = lean_ctor_get(x_75, 1); +x_99 = lean_ctor_get(x_75, 2); +x_100 = lean_ctor_get(x_75, 3); +x_101 = lean_ctor_get(x_75, 5); +x_102 = lean_ctor_get(x_75, 6); +lean_inc(x_102); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_75); +x_103 = l_Lean_Kernel_enableDiag(x_97, x_67); +x_104 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_98); +lean_ctor_set(x_104, 2, x_99); +lean_ctor_set(x_104, 3, x_100); +lean_ctor_set(x_104, 4, x_48); +lean_ctor_set(x_104, 5, x_101); +lean_ctor_set(x_104, 6, x_102); +x_105 = lean_st_ref_set(x_64, x_104, x_76); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_108 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__2(x_3, x_67, x_4, x_48, x_107, x_62, x_64, x_106); +if (lean_obj_tag(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_object* x_116; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_st_ref_get(x_64, x_110); +lean_dec(x_64); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; +} else { + lean_dec_ref(x_111); + x_114 = lean_box(0); +} +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); +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); +x_7 = x_116; +goto block_20; +} +else +{ +lean_object* x_117; lean_object* x_118; +lean_dec(x_64); +x_117 = lean_ctor_get(x_108, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_108, 1); +lean_inc(x_118); +lean_dec(x_108); +x_21 = x_117; +x_22 = x_118; +goto block_42; +} +} +} +else +{ +lean_object* x_119; lean_object* x_120; +x_119 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_120 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__2(x_3, x_67, x_4, x_48, x_119, x_62, x_64, x_70); +if (lean_obj_tag(x_120) == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +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_st_ref_get(x_64, x_122); +lean_dec(x_64); +x_124 = !lean_is_exclusive(x_123); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; +x_125 = lean_ctor_get(x_123, 0); +x_126 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_126, 0, x_121); +lean_ctor_set(x_126, 1, x_125); +lean_ctor_set(x_123, 0, x_126); +x_7 = x_123; +goto block_20; +} +else +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_127 = lean_ctor_get(x_123, 0); +x_128 = lean_ctor_get(x_123, 1); +lean_inc(x_128); +lean_inc(x_127); +lean_dec(x_123); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_121); +lean_ctor_set(x_129, 1, x_127); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_129); +lean_ctor_set(x_130, 1, x_128); +x_7 = x_130; +goto block_20; +} +} +else +{ +lean_object* x_131; lean_object* x_132; +lean_dec(x_64); +x_131 = lean_ctor_get(x_120, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_120, 1); +lean_inc(x_132); +lean_dec(x_120); +x_21 = x_131; +x_22 = x_132; +goto block_42; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_Core_getAndEmptyMessageLog(x_1); -lean_dec(x_1); +x_2 = lean_alloc_closure((void*)(l_Lean_Core_instMetaEvalCoreM___rarg___boxed), 6, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 1); -lean_inc(x_4); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; +size_t x_10; size_t x_11; lean_object* x_12; +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_12 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__4(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_Core_instMonadLogCoreM___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___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_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_3); -return x_5; +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__5(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___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) { _start: { -lean_object* x_4; uint8_t x_5; -x_4 = lean_st_ref_get(x_2, x_3); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) +lean_object* x_6; +x_6 = l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3___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_Core_instMetaEvalCoreM___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_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; -x_6 = lean_ctor_get(x_4, 0); -x_7 = lean_ctor_get(x_6, 5); -lean_inc(x_7); +size_t x_9; size_t x_10; lean_object* x_11; +x_9 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_10 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__6(x_1, x_2, x_9, x_10, x_5, x_6, x_7, x_8); +lean_dec(x_7); lean_dec(x_6); -x_8 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_7); -x_9 = lean_box(x_8); -lean_ctor_set(x_4, 0, x_9); -return x_4; +lean_dec(x_2); +return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___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) { +_start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; -x_10 = lean_ctor_get(x_4, 0); -x_11 = lean_ctor_get(x_4, 1); -lean_inc(x_11); -lean_inc(x_10); +lean_object* x_6; +x_6 = l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2___lambda__1(x_1, x_2, x_3, x_4, x_5); lean_dec(x_4); -x_12 = lean_ctor_get(x_10, 5); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_12); -x_14 = lean_box(x_13); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_11); -return x_15; -} +lean_dec(x_3); +lean_dec(x_2); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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_5; -x_5 = !lean_is_exclusive(x_1); -if (x_5 == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_6 = lean_ctor_get(x_1, 4); -x_7 = lean_ctor_get(x_2, 6); -x_8 = lean_ctor_get(x_2, 7); -lean_inc(x_8); -lean_inc(x_7); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_7); -lean_ctor_set(x_9, 1, x_8); -x_10 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_6); -lean_ctor_set(x_1, 4, x_10); -x_11 = lean_st_ref_take(x_3, x_4); -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) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_3); +lean_dec(x_3); +x_10 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_5); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___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: { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_12, 5); -x_16 = l_Lean_PersistentArray_push___rarg(x_15, x_1); -lean_ctor_set(x_12, 5, x_16); -x_17 = lean_st_ref_set(x_3, x_12, x_13); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_2); +lean_dec(x_2); +x_10 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__2(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_Core_instMetaEvalCoreM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 0); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_17, 0, x_20); -return x_17; +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_5); +lean_dec(x_5); +x_8 = l_Lean_Core_instMetaEvalCoreM___rarg(x_1, x_2, x_3, x_4, x_7, x_6); +return x_8; } -else +} +static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1() { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_17, 1); -lean_inc(x_21); -lean_dec(x_17); -x_22 = lean_box(0); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -return x_23; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_maxRecDepthErrorMessage; +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_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2() { +_start: { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; 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_24 = lean_ctor_get(x_12, 0); +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___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_Lean_Core_withIncRecDepth___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_5); +x_7 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed), 4, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___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) { +_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; uint8_t x_18; uint8_t x_19; uint8_t x_20; +lean_dec(x_2); +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_4, 1); +lean_inc(x_8); +x_9 = lean_ctor_get(x_4, 2); +lean_inc(x_9); +x_10 = lean_ctor_get(x_4, 3); +lean_inc(x_10); +x_11 = lean_ctor_get(x_4, 4); +lean_inc(x_11); +x_12 = lean_ctor_get(x_4, 5); +lean_inc(x_12); +x_13 = lean_ctor_get(x_4, 6); +lean_inc(x_13); +x_14 = lean_ctor_get(x_4, 7); +lean_inc(x_14); +x_15 = lean_ctor_get(x_4, 8); +lean_inc(x_15); +x_16 = lean_ctor_get(x_4, 9); +lean_inc(x_16); +x_17 = lean_ctor_get(x_4, 10); +lean_inc(x_17); +x_18 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); +x_19 = lean_ctor_get_uint8(x_4, sizeof(void*)*11 + 1); +x_20 = lean_nat_dec_eq(x_10, x_11); +if (x_20 == 0) +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_4); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_22 = lean_ctor_get(x_4, 10); +lean_dec(x_22); +x_23 = lean_ctor_get(x_4, 9); +lean_dec(x_23); +x_24 = lean_ctor_get(x_4, 8); +lean_dec(x_24); +x_25 = lean_ctor_get(x_4, 7); +lean_dec(x_25); +x_26 = lean_ctor_get(x_4, 6); +lean_dec(x_26); +x_27 = lean_ctor_get(x_4, 5); +lean_dec(x_27); +x_28 = lean_ctor_get(x_4, 4); +lean_dec(x_28); +x_29 = lean_ctor_get(x_4, 3); +lean_dec(x_29); +x_30 = lean_ctor_get(x_4, 2); +lean_dec(x_30); +x_31 = lean_ctor_get(x_4, 1); +lean_dec(x_31); +x_32 = lean_ctor_get(x_4, 0); +lean_dec(x_32); +x_33 = lean_unsigned_to_nat(1u); +x_34 = lean_nat_add(x_10, x_33); +lean_dec(x_10); +lean_ctor_set(x_4, 3, x_34); +x_35 = lean_apply_5(x_3, lean_box(0), x_1, x_4, x_5, x_6); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_4); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_add(x_10, x_36); +lean_dec(x_10); +x_38 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_38, 0, x_7); +lean_ctor_set(x_38, 1, x_8); +lean_ctor_set(x_38, 2, x_9); +lean_ctor_set(x_38, 3, x_37); +lean_ctor_set(x_38, 4, x_11); +lean_ctor_set(x_38, 5, x_12); +lean_ctor_set(x_38, 6, x_13); +lean_ctor_set(x_38, 7, x_14); +lean_ctor_set(x_38, 8, x_15); +lean_ctor_set(x_38, 9, x_16); +lean_ctor_set(x_38, 10, x_17); +lean_ctor_set_uint8(x_38, sizeof(void*)*11, x_18); +lean_ctor_set_uint8(x_38, sizeof(void*)*11 + 1, x_19); +x_39 = lean_apply_5(x_3, lean_box(0), x_1, x_38, x_5, x_6); +return x_39; +} +} +else +{ +lean_object* x_40; +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_1); +x_40 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg(x_12, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_40; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +lean_inc(x_2); +x_5 = lean_alloc_closure((void*)(l_Lean_Core_withIncRecDepth___rarg___lambda__1), 6, 2); +lean_closure_set(x_5, 0, x_3); +lean_closure_set(x_5, 1, x_2); +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +x_7 = lean_apply_2(x_6, lean_box(0), x_5); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_apply_1(x_8, lean_box(0)); +x_10 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Core_withIncRecDepth___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1(x_1, x_2, x_3); +lean_dec(x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("elaboration interrupted", 23); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_checkInterrupted___rarg___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_Core_checkInterrupted___rarg___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Core_checkInterrupted___rarg___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Core_checkInterrupted___rarg___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Core_checkInterrupted___rarg___closed__3; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; uint8_t x_4; +x_2 = lean_io_check_canceled(x_1); +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +x_4 = lean_unbox(x_3); +lean_dec(x_3); +if (x_4 == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_2); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_2, 0); +lean_dec(x_6); +x_7 = lean_box(0); +lean_ctor_set(x_2, 0, x_7); +return x_2; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +} +else +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_2); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_2, 0); +lean_dec(x_12); +x_13 = l_Lean_Core_checkInterrupted___rarg___closed__4; +lean_ctor_set_tag(x_2, 1); +lean_ctor_set(x_2, 0, x_13); +return x_2; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_2, 1); +lean_inc(x_14); +lean_dec(x_2); +x_15 = l_Lean_Core_checkInterrupted___rarg___closed__4; +x_16 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Core_checkInterrupted___rarg), 1, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Core_checkInterrupted(x_1, x_2); +lean_dec(x_2); +lean_dec(x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(deterministic) timeout at `", 28); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("`, maximum number of heartbeats (", 33); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(") has been reached\nuse `set_option ", 35); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ` to set the limit\nuse `set_option ", 41); +return x_1; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; +x_1 = l_Lean_diagnostics; +x_2 = lean_ctor_get(x_1, 0); +lean_inc(x_2); +x_3 = 1; +x_4 = l_Lean_Name_toString(x_2, x_3); +return x_4; +} +} +static lean_object* _init_l_Lean_Core_throwMaxHeartbeat___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" true` to get diagnostic information", 36); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat(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; lean_object* x_9; lean_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; +x_7 = 1; +x_8 = l_Lean_Name_toString(x_1, x_7); +x_9 = l_Lean_Core_throwMaxHeartbeat___closed__1; +x_10 = lean_string_append(x_9, x_8); +lean_dec(x_8); +x_11 = l_Lean_Core_throwMaxHeartbeat___closed__2; +x_12 = lean_string_append(x_10, x_11); +x_13 = lean_unsigned_to_nat(1000u); +x_14 = lean_nat_div(x_3, x_13); +x_15 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); +x_16 = lean_string_append(x_12, x_15); +lean_dec(x_15); +x_17 = l_Lean_Core_throwMaxHeartbeat___closed__3; +x_18 = lean_string_append(x_16, x_17); +x_19 = l_Lean_Name_toString(x_2, x_7); +x_20 = lean_string_append(x_18, x_19); +lean_dec(x_19); +x_21 = l_Lean_Core_throwMaxHeartbeat___closed__4; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Lean_Core_throwMaxHeartbeat___closed__5; +x_24 = lean_string_append(x_22, x_23); +x_25 = l_Lean_Core_throwMaxHeartbeat___closed__6; +x_26 = lean_string_append(x_24, x_25); +x_27 = lean_ctor_get(x_4, 5); +x_28 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_28, 0, x_26); +x_29 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_29, 0, x_28); +lean_inc(x_27); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_27); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_6); +return x_31; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat___boxed(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_Core_throwMaxHeartbeat(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_nat_dec_eq(x_3, x_7); +if (x_8 == 0) +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_io_get_num_heartbeats(x_6); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_4, 8); +x_14 = lean_nat_sub(x_11, x_13); +lean_dec(x_11); +x_15 = lean_nat_dec_lt(x_3, x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_box(0); +lean_ctor_set(x_9, 0, x_16); +return x_9; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_free_object(x_9); +x_17 = lean_box(0); +x_18 = l_Lean_Name_str___override(x_17, x_1); +x_19 = l_Lean_Core_throwMaxHeartbeat(x_18, x_2, x_3, x_4, x_5, x_12); +return x_19; +} +} +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_9, 0); +x_21 = lean_ctor_get(x_9, 1); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_9); +x_22 = lean_ctor_get(x_4, 8); +x_23 = lean_nat_sub(x_20, x_22); +lean_dec(x_20); +x_24 = lean_nat_dec_lt(x_3, x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; +lean_dec(x_2); +lean_dec(x_1); +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_21); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_box(0); +x_28 = l_Lean_Name_str___override(x_27, x_1); +x_29 = l_Lean_Core_throwMaxHeartbeat(x_28, x_2, x_3, x_4, x_5, x_21); +return x_29; +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_2); +lean_dec(x_1); +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_6); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore___boxed(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_Core_checkMaxHeartbeatsCore(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_2, 9); +x_6 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__2; +x_7 = l_Lean_Core_checkMaxHeartbeatsCore(x_1, x_6, x_5, x_2, x_3, x_4); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats___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_Core_checkMaxHeartbeats(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkSystem(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_io_check_canceled(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_unbox(x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = l_Lean_Core_checkMaxHeartbeats(x_1, x_2, x_3, x_8); +return x_9; +} +else +{ +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_5); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_5, 0); +lean_dec(x_11); +x_12 = l_Lean_Core_checkInterrupted___rarg___closed__4; +lean_ctor_set_tag(x_5, 1); +lean_ctor_set(x_5, 0, x_12); +return x_5; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_5, 1); +lean_inc(x_13); +lean_dec(x_5); +x_14 = l_Lean_Core_checkInterrupted___rarg___closed__4; +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_checkSystem___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_Core_checkSystem(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_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_io_get_num_heartbeats(x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = !lean_is_exclusive(x_2); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_2, 8); +lean_dec(x_9); +lean_ctor_set(x_2, 8, x_6); +x_10 = lean_apply_3(x_1, x_2, x_3, x_7); +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; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 1); +x_13 = lean_ctor_get(x_2, 2); +x_14 = lean_ctor_get(x_2, 3); +x_15 = lean_ctor_get(x_2, 4); +x_16 = lean_ctor_get(x_2, 5); +x_17 = lean_ctor_get(x_2, 6); +x_18 = lean_ctor_get(x_2, 7); +x_19 = lean_ctor_get(x_2, 9); +x_20 = lean_ctor_get(x_2, 10); +x_21 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); +x_22 = lean_ctor_get_uint8(x_2, sizeof(void*)*11 + 1); +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_12); +lean_inc(x_11); +lean_dec(x_2); +x_23 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_23, 0, x_11); +lean_ctor_set(x_23, 1, x_12); +lean_ctor_set(x_23, 2, x_13); +lean_ctor_set(x_23, 3, x_14); +lean_ctor_set(x_23, 4, x_15); +lean_ctor_set(x_23, 5, x_16); +lean_ctor_set(x_23, 6, x_17); +lean_ctor_set(x_23, 7, x_18); +lean_ctor_set(x_23, 8, x_6); +lean_ctor_set(x_23, 9, x_19); +lean_ctor_set(x_23, 10, x_20); +lean_ctor_set_uint8(x_23, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_23, sizeof(void*)*11 + 1, x_22); +x_24 = lean_apply_3(x_1, x_23, x_3, x_7); +return x_24; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___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_6; lean_object* x_7; +x_6 = lean_apply_2(x_2, lean_box(0), x_1); +x_7 = l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(x_6, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_alloc_closure((void*)(l_Lean_Core_withCurrHeartbeats___rarg___lambda__1), 5, 1); +lean_closure_set(x_5, 0, x_3); +x_6 = lean_ctor_get(x_2, 0); +lean_inc(x_6); +x_7 = lean_apply_2(x_6, lean_box(0), x_5); +x_8 = lean_ctor_get(x_2, 1); +lean_inc(x_8); +lean_dec(x_2); +x_9 = lean_apply_1(x_8, lean_box(0)); +x_10 = lean_apply_4(x_4, lean_box(0), lean_box(0), x_7, x_9); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Core_withCurrHeartbeats___rarg), 3, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_st_ref_take(x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_9 = lean_ctor_get(x_6, 5); +lean_dec(x_9); +lean_ctor_set(x_6, 5, x_1); +x_10 = lean_st_ref_set(x_3, x_6, x_7); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_dec(x_12); +x_13 = lean_box(0); +lean_ctor_set(x_10, 0, x_13); +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_10, 1); +lean_inc(x_14); +lean_dec(x_10); +x_15 = lean_box(0); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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_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, 6); +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_dec(x_6); +x_23 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_23, 0, x_17); +lean_ctor_set(x_23, 1, x_18); +lean_ctor_set(x_23, 2, x_19); +lean_ctor_set(x_23, 3, x_20); +lean_ctor_set(x_23, 4, x_21); +lean_ctor_set(x_23, 5, x_1); +lean_ctor_set(x_23, 6, x_22); +x_24 = lean_st_ref_set(x_3, x_23, x_7); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +if (lean_is_exclusive(x_24)) { + lean_ctor_release(x_24, 0); + lean_ctor_release(x_24, 1); + x_26 = x_24; +} else { + lean_dec_ref(x_24); + x_26 = lean_box(0); +} +x_27 = lean_box(0); +if (lean_is_scalar(x_26)) { + x_28 = lean_alloc_ctor(0, 2, 0); +} else { + x_28 = x_26; +} +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_25); +return x_28; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog___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_Core_setMessageLog(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = l_Lean_Core_State_traceState___default___closed__3; +x_5 = l_Lean_Core_setMessageLog(x_4, x_1, x_2, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Core_resetMessageLog(x_1, x_2, x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___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; +x_5 = lean_ctor_get(x_3, 0); +x_6 = lean_ctor_get(x_5, 5); +lean_inc(x_6); +lean_dec(x_5); +lean_ctor_set(x_3, 0, x_6); +return x_3; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_7 = lean_ctor_get(x_3, 0); +x_8 = lean_ctor_get(x_3, 1); +lean_inc(x_8); +lean_inc(x_7); +lean_dec(x_3); +x_9 = lean_ctor_get(x_7, 5); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Core_getMessageLog___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Core_getMessageLog___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Core_getMessageLog(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___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; uint8_t x_7; +x_3 = lean_st_ref_take(x_1, x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_st_ref_set(x_1, x_4, x_5); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_dec(x_8); +x_9 = l_Lean_Core_State_traceState___default___closed__3; +lean_ctor_set(x_6, 0, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_6, 1); +lean_inc(x_10); +lean_dec(x_6); +x_11 = l_Lean_Core_State_traceState___default___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_10); +return x_12; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Core_getAndEmptyMessageLog___rarg___boxed), 2, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_Core_getAndEmptyMessageLog___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Core_getAndEmptyMessageLog(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 1); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__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_st_ref_get(x_2, x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_6, 5); +lean_inc(x_7); +lean_dec(x_6); +x_8 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_7); +x_9 = lean_box(x_8); +lean_ctor_set(x_4, 0, x_9); +return x_4; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; +x_10 = lean_ctor_get(x_4, 0); +x_11 = lean_ctor_get(x_4, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_4); +x_12 = lean_ctor_get(x_10, 5); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_12); +x_14 = lean_box(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_11); +return x_15; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_1); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_6 = lean_ctor_get(x_1, 4); +x_7 = lean_ctor_get(x_2, 6); +x_8 = lean_ctor_get(x_2, 7); +lean_inc(x_8); +lean_inc(x_7); +x_9 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_9, 0, x_7); +lean_ctor_set(x_9, 1, x_8); +x_10 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_6); +lean_ctor_set(x_1, 4, x_10); +x_11 = lean_st_ref_take(x_3, x_4); +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; uint8_t x_18; +x_15 = lean_ctor_get(x_12, 5); +x_16 = l_Lean_PersistentArray_push___rarg(x_15, x_1); +lean_ctor_set(x_12, 5, x_16); +x_17 = lean_st_ref_set(x_3, x_12, x_13); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_dec(x_19); +x_20 = lean_box(0); +lean_ctor_set(x_17, 0, x_20); +return x_17; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 1); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_box(0); +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 +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; 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_24 = lean_ctor_get(x_12, 0); x_25 = lean_ctor_get(x_12, 1); x_26 = lean_ctor_get(x_12, 2); x_27 = lean_ctor_get(x_12, 3); @@ -10589,119 +12468,7 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_2; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_io_mono_nanos_now(x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_apply_3(x_1, x_2, x_3, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_io_mono_nanos_now(x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; lean_object* x_15; double x_16; double x_17; double x_18; double x_19; double x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_13 = lean_ctor_get(x_11, 0); -x_14 = 0; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Float_ofScientific(x_6, x_14, x_15); -lean_dec(x_6); -x_17 = l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10___closed__1; -x_18 = lean_float_div(x_16, x_17); -x_19 = l_Float_ofScientific(x_13, x_14, x_15); -lean_dec(x_13); -x_20 = lean_float_div(x_19, x_17); -x_21 = lean_box_float(x_18); -x_22 = lean_box_float(x_20); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_9); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_11, 0, x_24); -return x_11; -} -else -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_11); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Float_ofScientific(x_6, x_27, x_28); -lean_dec(x_6); -x_30 = l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10___closed__1; -x_31 = lean_float_div(x_29, x_30); -x_32 = l_Float_ofScientific(x_25, x_27, x_28); -lean_dec(x_25); -x_33 = lean_float_div(x_32, x_30); -x_34 = lean_box_float(x_31); -x_35 = lean_box_float(x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_9); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_26); -return x_38; -} -} -else -{ -uint8_t x_39; -lean_dec(x_6); -x_39 = !lean_is_exclusive(x_8); -if (x_39 == 0) -{ -return x_8; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_8, 0); -x_41 = lean_ctor_get(x_8, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_8); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; -} -} -} -} -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, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___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) { _start: { uint8_t x_8; @@ -10970,7 +12737,7 @@ return x_97; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -10995,88 +12762,22 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_apply_3(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_5, 0, x_8); -return x_5; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; -} -} -else -{ -uint8_t x_13; -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set_tag(x_5, 0); -lean_ctor_set(x_5, 0, x_15); -return x_5; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_5, 0); -x_17 = lean_ctor_get(x_5, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_5); -x_18 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_18, 0, x_16); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___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, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_inc(x_7); -x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__11(x_1, x_5, x_2, x_3, x_7, x_8, x_9); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10(x_1, x_5, x_2, x_3, x_7, x_8, x_9); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__12(x_4, x_7, x_8, x_11); +x_12 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(x_4, x_7, x_8, x_11); lean_dec(x_7); lean_dec(x_4); return x_12; } } -static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -11086,7 +12787,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -11094,11 +12795,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { double x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1; +x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_16 = lean_alloc_ctor(0, 2, 17); @@ -11107,7 +12808,7 @@ lean_ctor_set(x_16, 1, x_3); lean_ctor_set_float(x_16, sizeof(void*)*2, x_15); lean_ctor_set_float(x_16, sizeof(void*)*2 + 8, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 16, x_2); -x_17 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; +x_17 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2; x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_17); lean_dec(x_7); if (x_18 == 0) @@ -11118,7 +12819,7 @@ lean_object* x_19; lean_object* x_20; lean_dec(x_3); lean_dec(x_1); x_19 = lean_box(0); -x_20 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); +x_20 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); return x_20; } else @@ -11132,7 +12833,7 @@ lean_ctor_set_float(x_21, sizeof(void*)*2, x_9); lean_ctor_set_float(x_21, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 16, x_2); x_22 = lean_box(0); -x_23 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); +x_23 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); return x_23; } } @@ -11147,12 +12848,12 @@ lean_ctor_set_float(x_24, sizeof(void*)*2, x_9); lean_ctor_set_float(x_24, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_24, sizeof(void*)*2 + 16, x_2); x_25 = lean_box(0); -x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); +x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); return x_26; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -11160,16 +12861,16 @@ x_1 = lean_mk_string_from_bytes("", 9); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_compileDecl___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 uint8_t _init_l_Lean_ImportM_runCoreM___rarg___closed__3() { _start: { -lean_object* x_8; -x_8 = l_Lean_compileDecl___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); -return x_8; +lean_object* x_1; lean_object* x_2; uint8_t x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_7 = lean_st_ref_get(x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_ctor_get(x_8, 0); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_compile_decls(x_10, x_1, x_2); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -if (lean_obj_tag(x_12) == 12) +lean_object* x_4; lean_object* x_17; lean_object* x_18; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; 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; uint8_t x_69; +x_39 = lean_box(0); +x_40 = lean_ctor_get(x_2, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_2, 1); +lean_inc(x_41); +lean_dec(x_2); +x_42 = l_Lean_Core_State_nextMacroScope___default___closed__1; +x_43 = l_Lean_Core_State_ngen___default___closed__3; +x_44 = l_Lean_Core_State_traceState___default___closed__3; +x_45 = l_Lean_Core_instInhabitedCache___closed__1; +x_46 = l_Lean_Core_instMetaEvalCoreM___rarg___closed__3; +x_47 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_47, 0, x_40); +lean_ctor_set(x_47, 1, x_42); +lean_ctor_set(x_47, 2, x_43); +lean_ctor_set(x_47, 3, x_44); +lean_ctor_set(x_47, 4, x_45); +lean_ctor_set(x_47, 5, x_44); +lean_ctor_set(x_47, 6, x_46); +x_48 = lean_io_get_num_heartbeats(x_3); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_48, 1); +lean_inc(x_50); +lean_dec(x_48); +x_51 = l_Lean_ImportM_runCoreM___rarg___closed__2; +x_52 = l_Lean_Core_instMetaEvalCoreM___rarg___closed__2; +x_53 = lean_unsigned_to_nat(0u); +x_54 = lean_unsigned_to_nat(1000u); +x_55 = lean_box(0); +x_56 = lean_box(0); +x_57 = l_Lean_ImportM_runCoreM___rarg___closed__1; +x_58 = l_Lean_firstFrontendMacroScope; +x_59 = 0; +x_60 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_60, 0, x_51); +lean_ctor_set(x_60, 1, x_52); +lean_ctor_set(x_60, 2, x_39); +lean_ctor_set(x_60, 3, x_53); +lean_ctor_set(x_60, 4, x_54); +lean_ctor_set(x_60, 5, x_55); +lean_ctor_set(x_60, 6, x_56); +lean_ctor_set(x_60, 7, x_39); +lean_ctor_set(x_60, 8, x_49); +lean_ctor_set(x_60, 9, x_57); +lean_ctor_set(x_60, 10, x_58); +lean_ctor_set_uint8(x_60, sizeof(void*)*11, x_59); +lean_ctor_set_uint8(x_60, sizeof(void*)*11 + 1, x_59); +x_61 = lean_st_mk_ref(x_47, x_50); +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_st_ref_get(x_62, x_63); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_ctor_get(x_65, 0); +lean_inc(x_67); +lean_dec(x_65); +x_68 = l_Lean_Kernel_isDiagnosticsEnabled(x_67); +lean_dec(x_67); +if (x_68 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_14, 0, x_13); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -x_16 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_15, x_4, x_5, x_9); -lean_dec(x_4); -return x_16; -} -else +uint8_t x_133; +x_133 = l_Lean_ImportM_runCoreM___rarg___closed__3; +if (x_133 == 0) { -lean_object* x_17; -x_17 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_12, x_4, x_5, x_9); -return x_17; -} +uint8_t x_134; +x_134 = 1; +x_69 = x_134; +goto block_132; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_11, 0); -lean_inc(x_18); -lean_dec(x_11); -x_19 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_18, x_4, x_5, x_9); -lean_dec(x_4); -return x_19; -} -} -} -LEAN_EXPORT lean_object* l_Lean_compileDecls(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 2); -lean_inc(x_5); -x_6 = l_Lean_compileDecl___closed__1; -x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_box(0); -x_9 = l_Lean_compileDecls___lambda__1(x_5, x_1, x_8, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_1); -lean_dec(x_5); -return x_9; +x_69 = x_59; +goto block_132; } -else -{ -lean_object* x_10; -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_10 = lean_lcnf_compile_decls(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_compileDecls___lambda__1(x_5, x_1, x_11, x_2, x_3, x_12); -lean_dec(x_3); -lean_dec(x_11); -lean_dec(x_1); -lean_dec(x_5); -return x_13; } else { -uint8_t x_14; -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_14 = !lean_is_exclusive(x_10); -if (x_14 == 0) +uint8_t x_135; +x_135 = l_Lean_ImportM_runCoreM___rarg___closed__3; +if (x_135 == 0) { -return x_10; +x_69 = x_59; +goto block_132; } 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; -} -} -} +uint8_t x_136; +x_136 = 1; +x_69 = x_136; +goto block_132; } } -LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +block_16: { -lean_object* x_7; -x_7 = l_Lean_compileDecls___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +if (lean_obj_tag(x_4) == 0) { -lean_object* x_5; -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_5 = l_Lean_addDecl(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) { lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_compileDecl(x_1, x_2, x_3, x_6); -return x_7; -} -else -{ -uint8_t x_8; -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) -{ -return x_5; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +lean_ctor_set(x_4, 0, x_7); +return x_4; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); lean_inc(x_9); -lean_dec(x_5); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); +lean_inc(x_8); +lean_dec(x_4); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); return x_11; } } -} -} -LEAN_EXPORT uint8_t l_Lean_getDiag(lean_object* x_1) { -_start: +else { -lean_object* x_2; uint8_t x_3; -x_2 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4; -x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); -return x_3; -} +uint8_t x_12; +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) +{ +return x_4; } -LEAN_EXPORT lean_object* l_Lean_getDiag___boxed(lean_object* x_1) { -_start: +else { -uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_getDiag(x_1); -lean_dec(x_1); -x_3 = lean_box(x_2); -return x_3; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_isDiagnosticsEnabled(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get_uint8(x_1, sizeof(void*)*11 + 1); -x_5 = lean_box(x_4); -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_3); -return x_6; } +block_38: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_MessageData_toString(x_19, x_18); +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_alloc_ctor(18, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set_tag(x_20, 1); +lean_ctor_set(x_20, 0, x_23); +x_4 = x_20; +goto block_16; } -LEAN_EXPORT lean_object* l_Lean_isDiagnosticsEnabled___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Lean_isDiagnosticsEnabled(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; +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_alloc_ctor(18, 1, 0); +lean_ctor_set(x_26, 0, x_24); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_4 = x_27; +goto block_16; } } -static lean_object* _init_l_Lean_ImportM_runCoreM___rarg___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 9); -return x_1; -} +uint8_t x_28; +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) +{ +x_4 = x_20; +goto block_16; } -LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_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; lean_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_29; lean_object* x_30; lean_object* x_31; -x_4 = lean_ctor_get(x_2, 0); -x_5 = lean_ctor_get(x_2, 1); -x_6 = lean_box(0); -x_7 = l_Lean_Core_getMaxHeartbeats(x_5); -x_8 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4; -x_9 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_8); -x_10 = l_Lean_Core_State_nextMacroScope___default___closed__1; -x_11 = l_Lean_Core_State_ngen___default___closed__3; -x_12 = l_Lean_Core_State_traceState___default___closed__3; -x_13 = l_Lean_Core_instInhabitedCache___closed__1; -x_14 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5; -lean_inc(x_4); -x_15 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_15, 0, x_4); -lean_ctor_set(x_15, 1, x_10); -lean_ctor_set(x_15, 2, x_11); -lean_ctor_set(x_15, 3, x_12); -lean_ctor_set(x_15, 4, x_13); -lean_ctor_set(x_15, 5, x_12); -lean_ctor_set(x_15, 6, x_14); -x_16 = lean_io_get_num_heartbeats(x_3); -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_ImportM_runCoreM___rarg___closed__1; -x_20 = l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__2; -x_21 = lean_unsigned_to_nat(0u); -x_22 = lean_unsigned_to_nat(1000u); -x_23 = lean_box(0); -x_24 = lean_box(0); -x_25 = l_Lean_firstFrontendMacroScope; -x_26 = 0; -lean_inc(x_5); -x_27 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_27, 0, x_19); -lean_ctor_set(x_27, 1, x_20); -lean_ctor_set(x_27, 2, x_5); -lean_ctor_set(x_27, 3, x_21); -lean_ctor_set(x_27, 4, x_22); -lean_ctor_set(x_27, 5, x_23); -lean_ctor_set(x_27, 6, x_24); -lean_ctor_set(x_27, 7, x_6); -lean_ctor_set(x_27, 8, x_17); -lean_ctor_set(x_27, 9, x_7); -lean_ctor_set(x_27, 10, x_25); -lean_ctor_set_uint8(x_27, sizeof(void*)*11, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*11 + 1, x_9); -x_28 = lean_st_mk_ref(x_15, x_18); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); +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_dec(x_28); lean_inc(x_29); -x_31 = lean_apply_3(x_1, x_27, x_29, x_30); -if (lean_obj_tag(x_31) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_st_ref_get(x_29, x_33); -lean_dec(x_29); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) -{ -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_32); -return x_34; +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); +x_4 = x_31; +goto block_16; +} +} } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_37); -return x_38; +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_32 = lean_ctor_get(x_17, 0); +lean_inc(x_32); +lean_dec(x_17); +x_33 = l___private_Init_Data_Repr_0__Nat_reprFast(x_32); +x_34 = l_Lean_Core_CoreM_toIO___rarg___closed__1; +x_35 = lean_string_append(x_34, x_33); +lean_dec(x_33); +x_36 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_18); +x_4 = x_37; +goto block_16; } } -else +block_132: { -lean_object* x_39; -lean_dec(x_29); -x_39 = lean_ctor_get(x_31, 0); -lean_inc(x_39); -if (lean_obj_tag(x_39) == 0) +if (x_69 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_31, 1); -lean_inc(x_40); -lean_dec(x_31); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_MessageData_toString(x_41, x_40); -if (lean_obj_tag(x_42) == 0) +lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_70 = lean_st_ref_take(x_62, x_66); +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_is_exclusive(x_71); +if (x_73 == 0) { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_object* x_74; lean_object* x_75; uint8_t x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_74 = lean_ctor_get(x_71, 0); +x_75 = lean_ctor_get(x_71, 4); +lean_dec(x_75); +x_76 = l_Lean_ImportM_runCoreM___rarg___closed__3; +x_77 = l_Lean_Kernel_enableDiag(x_74, x_76); +lean_ctor_set(x_71, 4, x_45); +lean_ctor_set(x_71, 0, x_77); +x_78 = lean_st_ref_set(x_62, x_71, x_72); +x_79 = lean_ctor_get(x_78, 1); +lean_inc(x_79); +lean_dec(x_78); +x_80 = lean_box(0); +lean_inc(x_62); +x_81 = l_Lean_ImportM_runCoreM___rarg___lambda__1(x_39, x_76, x_41, x_1, x_45, x_80, x_60, x_62, x_79); +if (lean_obj_tag(x_81) == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set_tag(x_42, 1); -lean_ctor_set(x_42, 0, x_45); -return x_42; +lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = lean_st_ref_get(x_62, x_83); +lean_dec(x_62); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) +{ +lean_object* x_86; lean_object* x_87; +x_86 = lean_ctor_get(x_84, 0); +x_87 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_87, 0, x_82); +lean_ctor_set(x_87, 1, x_86); +lean_ctor_set(x_84, 0, x_87); +x_4 = x_84; +goto block_16; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_46 = lean_ctor_get(x_42, 0); -x_47 = lean_ctor_get(x_42, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_42); -x_48 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_48, 0, x_46); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -return x_49; +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +x_88 = lean_ctor_get(x_84, 0); +x_89 = lean_ctor_get(x_84, 1); +lean_inc(x_89); +lean_inc(x_88); +lean_dec(x_84); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_82); +lean_ctor_set(x_90, 1, x_88); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_89); +x_4 = x_91; +goto block_16; } } else { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_42); -if (x_50 == 0) +lean_object* x_92; lean_object* x_93; +lean_dec(x_62); +x_92 = lean_ctor_get(x_81, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_81, 1); +lean_inc(x_93); +lean_dec(x_81); +x_17 = x_92; +x_18 = x_93; +goto block_38; +} +} +else { -return x_42; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_94 = lean_ctor_get(x_71, 0); +x_95 = lean_ctor_get(x_71, 1); +x_96 = lean_ctor_get(x_71, 2); +x_97 = lean_ctor_get(x_71, 3); +x_98 = lean_ctor_get(x_71, 5); +x_99 = lean_ctor_get(x_71, 6); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); +lean_inc(x_96); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_71); +x_100 = l_Lean_ImportM_runCoreM___rarg___closed__3; +x_101 = l_Lean_Kernel_enableDiag(x_94, x_100); +x_102 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_95); +lean_ctor_set(x_102, 2, x_96); +lean_ctor_set(x_102, 3, x_97); +lean_ctor_set(x_102, 4, x_45); +lean_ctor_set(x_102, 5, x_98); +lean_ctor_set(x_102, 6, x_99); +x_103 = lean_st_ref_set(x_62, x_102, x_72); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +lean_dec(x_103); +x_105 = lean_box(0); +lean_inc(x_62); +x_106 = l_Lean_ImportM_runCoreM___rarg___lambda__1(x_39, x_100, x_41, x_1, x_45, x_105, x_60, x_62, x_104); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_st_ref_get(x_62, x_108); +lean_dec(x_62); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_112 = x_109; +} else { + lean_dec_ref(x_109); + x_112 = lean_box(0); +} +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_107); +lean_ctor_set(x_113, 1, x_110); +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_112; +} +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_111); +x_4 = x_114; +goto block_16; } 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_115; lean_object* x_116; +lean_dec(x_62); +x_115 = lean_ctor_get(x_106, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_106, 1); +lean_inc(x_116); +lean_dec(x_106); +x_17 = x_115; +x_18 = x_116; +goto block_38; } } } else { -uint8_t x_54; -x_54 = !lean_is_exclusive(x_31); -if (x_54 == 0) +uint8_t x_117; lean_object* x_118; lean_object* x_119; +x_117 = l_Lean_ImportM_runCoreM___rarg___closed__3; +x_118 = lean_box(0); +lean_inc(x_62); +x_119 = l_Lean_ImportM_runCoreM___rarg___lambda__1(x_39, x_117, x_41, x_1, x_45, x_118, x_60, x_62, x_66); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_55 = lean_ctor_get(x_31, 0); -lean_dec(x_55); -x_56 = lean_ctor_get(x_39, 0); -lean_inc(x_56); -lean_dec(x_39); -x_57 = l___private_Init_Data_Repr_0__Nat_reprFast(x_56); -x_58 = l_Lean_Core_CoreM_toIO___rarg___closed__1; -x_59 = lean_string_append(x_58, x_57); -lean_dec(x_57); -x_60 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_31, 0, x_60); -return x_31; +lean_object* x_120; lean_object* x_121; lean_object* x_122; uint8_t x_123; +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_122 = lean_st_ref_get(x_62, x_121); +lean_dec(x_62); +x_123 = !lean_is_exclusive(x_122); +if (x_123 == 0) +{ +lean_object* x_124; lean_object* x_125; +x_124 = lean_ctor_get(x_122, 0); +x_125 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_125, 0, x_120); +lean_ctor_set(x_125, 1, x_124); +lean_ctor_set(x_122, 0, x_125); +x_4 = x_122; +goto block_16; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_61 = lean_ctor_get(x_31, 1); -lean_inc(x_61); -lean_dec(x_31); -x_62 = lean_ctor_get(x_39, 0); -lean_inc(x_62); -lean_dec(x_39); -x_63 = l___private_Init_Data_Repr_0__Nat_reprFast(x_62); -x_64 = l_Lean_Core_CoreM_toIO___rarg___closed__1; -x_65 = lean_string_append(x_64, x_63); -lean_dec(x_63); -x_66 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_61); -return x_67; +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_126 = lean_ctor_get(x_122, 0); +x_127 = lean_ctor_get(x_122, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_122); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_120); +lean_ctor_set(x_128, 1, x_126); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_128); +lean_ctor_set(x_129, 1, x_127); +x_4 = x_129; +goto block_16; +} +} +else +{ +lean_object* x_130; lean_object* x_131; +lean_dec(x_62); +x_130 = lean_ctor_get(x_119, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_119, 1); +lean_inc(x_131); +lean_dec(x_119); +x_17 = x_130; +x_18 = x_131; +goto block_38; } } } @@ -15687,17 +18634,19 @@ LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ImportM_runCoreM___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_ImportM_runCoreM___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___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) { _start: { -lean_object* x_4; -x_4 = l_Lean_ImportM_runCoreM___rarg(x_1, x_2, x_3); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); lean_dec(x_2); -return x_4; +x_11 = l_Lean_ImportM_runCoreM___rarg___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_11; } } LEAN_EXPORT uint8_t l_Lean_Exception_isRuntime(lean_object* x_1) { @@ -16312,12 +19261,10 @@ l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__5 = _init_l_Lean_initFn____x4 lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__5); l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__6(); lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__6); -l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__7 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__7(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__7); if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_6_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; -l_Lean_diag = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_diag); +l_Lean_diagnostics = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_diagnostics); lean_dec_ref(res); }l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__1(); lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__1); @@ -16331,8 +19278,8 @@ l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__5 = _init_l_Lean_initFn____x lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__5); if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_41_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; -l_Lean_diag_threshold = lean_io_result_get_value(res); -lean_mark_persistent(l_Lean_diag_threshold); +l_Lean_diagnostics_threshold = lean_io_result_get_value(res); +lean_mark_persistent(l_Lean_diagnostics_threshold); lean_dec_ref(res); }l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__1(); lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__1); @@ -16484,6 +19431,10 @@ l_Lean_Core_instMonadEnvCoreM___closed__3 = _init_l_Lean_Core_instMonadEnvCoreM_ lean_mark_persistent(l_Lean_Core_instMonadEnvCoreM___closed__3); l_Lean_Core_instMonadEnvCoreM = _init_l_Lean_Core_instMonadEnvCoreM(); lean_mark_persistent(l_Lean_Core_instMonadEnvCoreM); +l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1 = _init_l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1); +l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1 = _init_l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1(); +lean_mark_persistent(l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1); l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__1 = _init_l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__1(); lean_mark_persistent(l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__1); l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__2 = _init_l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__2(); @@ -16560,18 +19511,14 @@ l_Lean_Core_instMonadTraceCoreM = _init_l_Lean_Core_instMonadTraceCoreM(); lean_mark_persistent(l_Lean_Core_instMonadTraceCoreM); l_Lean_Core_CoreM_toIO___rarg___closed__1 = _init_l_Lean_Core_CoreM_toIO___rarg___closed__1(); lean_mark_persistent(l_Lean_Core_CoreM_toIO___rarg___closed__1); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__1 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__1); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__2 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__2); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__3 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__3); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__4); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__5); -l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__6 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___closed__6); +l_Lean_Core_instMetaEvalCoreM___rarg___closed__1 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__1(); +lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___closed__1); +l_Lean_Core_instMetaEvalCoreM___rarg___closed__2 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__2(); +lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___closed__2); +l_Lean_Core_instMetaEvalCoreM___rarg___closed__3 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__3(); +lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___closed__3); +l_Lean_Core_instMetaEvalCoreM___rarg___closed__4 = _init_l_Lean_Core_instMetaEvalCoreM___rarg___closed__4(); +lean_mark_persistent(l_Lean_Core_instMetaEvalCoreM___rarg___closed__4); l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1); l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2(); @@ -16592,6 +19539,10 @@ l_Lean_Core_throwMaxHeartbeat___closed__3 = _init_l_Lean_Core_throwMaxHeartbeat_ lean_mark_persistent(l_Lean_Core_throwMaxHeartbeat___closed__3); l_Lean_Core_throwMaxHeartbeat___closed__4 = _init_l_Lean_Core_throwMaxHeartbeat___closed__4(); lean_mark_persistent(l_Lean_Core_throwMaxHeartbeat___closed__4); +l_Lean_Core_throwMaxHeartbeat___closed__5 = _init_l_Lean_Core_throwMaxHeartbeat___closed__5(); +lean_mark_persistent(l_Lean_Core_throwMaxHeartbeat___closed__5); +l_Lean_Core_throwMaxHeartbeat___closed__6 = _init_l_Lean_Core_throwMaxHeartbeat___closed__6(); +lean_mark_persistent(l_Lean_Core_throwMaxHeartbeat___closed__6); l_Lean_Core_instMonadLogCoreM___closed__1 = _init_l_Lean_Core_instMonadLogCoreM___closed__1(); lean_mark_persistent(l_Lean_Core_instMonadLogCoreM___closed__1); l_Lean_Core_instMonadLogCoreM___closed__2 = _init_l_Lean_Core_instMonadLogCoreM___closed__2(); @@ -16614,16 +19565,21 @@ l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1 = _init_l_Lean_logAt___at_L lean_mark_persistent(l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1); l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1(); lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1); -l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10___closed__1 = _init_l_Lean_withStartStopSeconds___at_Lean_addDecl___spec__10___closed__1(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2); l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1); l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2); l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1); l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2(); lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__5___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__5___closed__1); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5(); l_Lean_addDecl___lambda__1___closed__1 = _init_l_Lean_addDecl___lambda__1___closed__1(); lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__1); l_Lean_addDecl___lambda__1___closed__2 = _init_l_Lean_addDecl___lambda__1___closed__2(); @@ -16706,19 +19662,19 @@ l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rar lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3); l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4 = _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4(); lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__1); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__2 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__2); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__3 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__3); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__4 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__4); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__5 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__5); -l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3748____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_3748_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5); +l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_3898_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_compiler_enableNew = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_compiler_enableNew); @@ -16733,6 +19689,9 @@ l_Lean_compileDecl___closed__1 = _init_l_Lean_compileDecl___closed__1(); lean_mark_persistent(l_Lean_compileDecl___closed__1); l_Lean_ImportM_runCoreM___rarg___closed__1 = _init_l_Lean_ImportM_runCoreM___rarg___closed__1(); lean_mark_persistent(l_Lean_ImportM_runCoreM___rarg___closed__1); +l_Lean_ImportM_runCoreM___rarg___closed__2 = _init_l_Lean_ImportM_runCoreM___rarg___closed__2(); +lean_mark_persistent(l_Lean_ImportM_runCoreM___rarg___closed__2); +l_Lean_ImportM_runCoreM___rarg___closed__3 = _init_l_Lean_ImportM_runCoreM___rarg___closed__3(); l_Lean_instMonadExceptOfExceptionCoreM___closed__1 = _init_l_Lean_instMonadExceptOfExceptionCoreM___closed__1(); lean_mark_persistent(l_Lean_instMonadExceptOfExceptionCoreM___closed__1); l_Lean_instMonadExceptOfExceptionCoreM___closed__2 = _init_l_Lean_instMonadExceptOfExceptionCoreM___closed__2(); diff --git a/stage0/stdlib/Lean/Data/HashMap.c b/stage0/stdlib/Lean/Data/HashMap.c index 80ffe2f8bac4..ea608af7b61e 100644 --- a/stage0/stdlib/Lean/Data/HashMap.c +++ b/stage0/stdlib/Lean/Data/HashMap.c @@ -2630,7 +2630,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(212u); +x_3 = lean_unsigned_to_nat(213u); 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/Data/Lsp/Capabilities.c b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c index 126dca4e9a8b..e734db294101 100644 --- a/stage0/stdlib/Lean/Data/Lsp/Capabilities.c +++ b/stage0/stdlib/Lean/Data/Lsp/Capabilities.c @@ -26,7 +26,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonSer LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____spec__3___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonCompletionItemCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_44____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonServerCapabilities; -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__12; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1389____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonShowDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_370_(uint8_t); @@ -57,6 +56,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonT LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_935____spec__1___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__17; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__9; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8866_(lean_object*); LEAN_EXPORT uint8_t l_Lean_Lsp_ServerCapabilities_workspaceSymbolProvider___default; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__6; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); @@ -74,6 +74,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonS lean_object* l___private_Lean_Data_Lsp_TextSync_0__Lean_Lsp_fromJsonTextDocumentSyncOptions____x40_Lean_Data_Lsp_TextSync___hyg_1139_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCompletionItemCapabilities___closed__1; static lean_object* l_Lean_Lsp_instFromJsonShowDocumentClientCapabilities___closed__1; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650_(uint8_t); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1068____spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonServerCapabilities___closed__1; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__4; @@ -99,6 +100,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonW static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__51; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__68; LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_697____spec__1___boxed(lean_object*, lean_object*); +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonChangeAnnotationSupport____x40_Lean_Data_Lsp_Capabilities___hyg_600____closed__4; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__7; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonShowDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_402____closed__2; @@ -158,7 +160,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonW static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWindowClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_501____closed__5; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonChangeAnnotationSupport____x40_Lean_Data_Lsp_Capabilities___hyg_600_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_ClientCapabilities_window_x3f___default; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_736____closed__12; lean_object* l___private_Lean_Data_Lsp_CodeActions_0__Lean_Lsp_toJsonCodeActionClientCapabilities____x40_Lean_Data_Lsp_CodeActions___hyg_2280_(lean_object*); @@ -174,7 +175,6 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonC static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonTextDocumentClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_259____closed__11; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2259____spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582_(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647_(uint8_t); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__58; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__32; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_935____closed__1; @@ -294,6 +294,7 @@ static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonSer LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____spec__4(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__66; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonShowDocumentClientCapabilities; +lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1107____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonWorkspaceEditClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_736____spec__1___boxed(lean_object*, lean_object*); @@ -307,7 +308,6 @@ static lean_object* l_Lean_Lsp_instToJsonWorkspaceClientCapabilities___closed__1 LEAN_EXPORT uint8_t l_Lean_Lsp_ServerCapabilities_documentHighlightProvider___default; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__56; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonCompletionClientCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_118____boxed(lean_object*); -lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8863_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonCompletionItemCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_19____boxed(lean_object*); static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_fromJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1582____closed__45; static lean_object* l___private_Lean_Data_Lsp_Capabilities_0__Lean_Lsp_toJsonServerCapabilities____x40_Lean_Data_Lsp_Capabilities___hyg_1389____closed__14; @@ -4233,7 +4233,7 @@ lean_inc(x_4); lean_dec(x_2); x_5 = lean_unbox(x_4); lean_dec(x_4); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647_(x_5); +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650_(x_5); x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_1); lean_ctor_set(x_7, 1, x_6); @@ -4261,7 +4261,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8863_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8866_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -4755,7 +4755,7 @@ return x_4; else { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589_(x_3); lean_dec(x_3); if (lean_obj_tag(x_5) == 0) { @@ -4819,7 +4819,7 @@ return x_4; else { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727_(x_3); lean_dec(x_3); if (lean_obj_tag(x_5) == 0) { diff --git a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c index 07a630d34bd5..8e3f477bbb35 100644 --- a/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c +++ b/stage0/stdlib/Lean/Data/Lsp/LanguageFeatures.c @@ -13,1440 +13,1440 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_JsonNumber_fromNat(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__27; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokens; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182____spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableCallHierarchyItem; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66; lean_object* l_List_join___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams; static lean_object* l_Lean_Lsp_instToJsonCompletionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__17; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__51; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339_(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceContext; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__17; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__36; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; static lean_object* l_Lean_Lsp_instToJsonSymbolInformation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceContext; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078_(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat(uint8_t); static lean_object* l_Lean_Lsp_instToJsonCompletionItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHover; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__16; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9; static lean_object* l_Lean_Lsp_instFromJsonCompletionParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokenModifier; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43; static lean_object* l_Lean_Lsp_instToJsonFoldingRange___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__32; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3(lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CallHierarchyItem_detail_x3f___default; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__28; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5; +LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default; lean_object* l_Lean_Json_mkObj(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812_(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__47; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2(size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instHashableCallHierarchyItem___closed__1; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__37; LEAN_EXPORT uint8_t l_Lean_Lsp_RenameOptions_prepareProvider___default; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__35; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__33; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHoverParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796____boxed(lean_object*, lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97; static lean_object* l_Lean_Lsp_instBEqSymbolKind___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912_(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__22; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2646_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__4; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__45; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8866_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748____boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____boxed(lean_object*); static lean_object* l_Lean_Lsp_SymbolInformation_tags___default___closed__1; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127; static lean_object* l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9146_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedSymbolTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92; uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83; lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2803_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1380_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__45; static lean_object* l_Lean_Lsp_instToJsonReferenceParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__28; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__1(lean_object*, lean_object*); +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1(uint64_t, lean_object*, size_t, size_t, uint64_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10008_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767_(uint8_t); +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11; uint8_t lean_usize_dec_eq(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokens_resultId_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122; lean_object* l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2492_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_ofNat___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2963_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650_(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4; static lean_object* l_Lean_Lsp_instFromJsonReferenceParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; LEAN_EXPORT uint8_t l_Lean_Lsp_CompletionItemKind_ofNat(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_names; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instDecidableEqCompletionItemKind___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentSymbol___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__43; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8; static lean_object* l_Lean_Lsp_instFromJsonReferenceContext___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__22; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29; LEAN_EXPORT uint8_t l_Lean_Lsp_instDecidableEqCompletionItemKind(uint8_t, uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensRangeParams; lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokens; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8995_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag___boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_detail_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8; static lean_object* l_Lean_Lsp_instFromJsonHover___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3; uint64_t lean_string_hash(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlight_kind_x3f___default; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonHover; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14; -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9428_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_kind_x3f___default; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDeclarationParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1; static lean_object* l_Lean_Lsp_instBEqSymbolTag___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__45; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1; uint8_t lean_string_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__2; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionOptions_triggerCharacters_x3f___default; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2960_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2175_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokenType; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19; uint8_t l___private_Lean_Data_Json_Basic_0__Lean_Json_beq_x27(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1; uint64_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_674_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__40; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonHover___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__19; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonPrepareRenameParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6702_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSemanticTokenType; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64; -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____boxed(lean_object*); static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__21; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__49; static lean_object* l_Lean_Lsp_instToJsonCompletionOptions___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86; lean_object* l_Lean_Json_getStr_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__2; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764_(uint8_t); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1038_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____boxed(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonHoverParams; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__3; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55; size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__18; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__25; uint64_t l___private_Lean_Data_Json_Basic_0__Lean_Json_hash_x27(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toCtorIdx___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolInformation_containerName_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10011_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_sortText_x3f___default; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6547_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDeclarationParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_Hover_range_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9; static lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCompletionItem___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1383_(lean_object*); lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__36; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4; LEAN_EXPORT lean_object* l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1(size_t, size_t, lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5792_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22; lean_object* lean_nat_to_int(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4; lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1284____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586_(lean_object*); +LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__14; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2___boxed(lean_object*, lean_object*, 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toNat___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__45; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14; lean_object* l_Lean_Json_getObjValD(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolInformation_tags___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3543_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8998_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__23; -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14; +LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126; static lean_object* l_Lean_Lsp_instFromJsonSymbolInformation___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenModifier; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85; static lean_object* l_Lean_Lsp_instFromJsonSymbolTag___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647_(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3271_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__10; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__41; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645_(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonReferenceParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__39; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_documentation_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokenType___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CallHierarchyItem_tags_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__32; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCall; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__24; static lean_object* l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toCtorIdx(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_CallHierarchyItem_data_x3f___default; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__24; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11; static lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6; static lean_object* l_Lean_Lsp_instHashableSymbolTag___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132_(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22; static lean_object* l_Lean_Lsp_instFromJsonHoverParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSemanticTokensOptions; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759_(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342_(uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind(uint8_t); LEAN_EXPORT uint8_t l_Lean_Lsp_CompletionOptions_resolveProvider___default; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonPrepareRenameParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24; static lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__50; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18; static lean_object* l_Lean_Lsp_instHashableSymbolKind___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524_(lean_object*); LEAN_EXPORT 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*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyPrepareParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionList; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751_(uint8_t, uint8_t); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__25; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__37; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSemanticTokenType; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemKind___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqCallHierarchyItem; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCompletionItem; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__33; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__46; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3760_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol_go(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__21; static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2649_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__35; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__30; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2806_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRange_kind_x3f___default; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__29; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6267_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__18; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11; static lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__46; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102; static lean_object* l_Lean_Lsp_instToJsonReferenceContext___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__31; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRange; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonInsertReplaceEdit; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2(size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__9; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__40; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItemKind(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124; +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70; static lean_object* l_Lean_Lsp_instToJsonDefinitionParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSymbolTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instReprCompletionItemKind; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__5; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42; static lean_object* l_Lean_Lsp_instFromJsonRenameOptions___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionList; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4; lean_object* l_Lean_Json_parseTagged(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6; lean_object* lean_nat_abs(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentSymbolAux_children_x3f___default(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9837_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonTypeDefinitionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3268_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toNat(uint8_t); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8; static lean_object* l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionOptions; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647____boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__2(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6550_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; static lean_object* l_Lean_Lsp_instFromJsonSymbolTag___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9840_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__42; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2489_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3403_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8; LEAN_EXPORT lean_object* l_Lean_Lsp_DocumentHighlightKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__17; lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedSymbolKind; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDefinitionParams; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748_(uint8_t, uint8_t); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__39; LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams; -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonReferenceParams; static lean_object* l_Lean_Lsp_instToJsonSemanticTokens___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__25; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_names; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__34; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131; static lean_object* l_Lean_Lsp_instInhabitedCompletionItem___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolTag(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolInformation; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17; -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__1(uint64_t, lean_object*, size_t, size_t, uint64_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182____spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonCompletionList___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__26; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__15; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__31; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonWorkspaceSymbolParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8863_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__2; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolKind_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolKind; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3; static lean_object* l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_textEdit_x3f___default; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23; static lean_object* l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5; -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__23; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4; uint8_t lean_int_dec_lt(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54; static lean_object* l_Lean_Lsp_instToJsonHoverParams___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9; lean_object* l_Except_orElseLazy___rarg(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItem_data_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18; lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__52; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9425_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____boxed(lean_object*, lean_object*); lean_object* l_Lean_Json_getNat_x3f(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6; +LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instHashableSymbolTag; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDeclarationParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokensOptions; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__22; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__8; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078____boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__20; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__41; static lean_object* l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__48; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolResult(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5; static lean_object* l_Lean_Lsp_instHashableSemanticTokenType___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3400_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionOptions; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbol; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98; LEAN_EXPORT uint8_t l_Lean_Lsp_instInhabitedCompletionItemKind; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7; size_t lean_usize_add(size_t, size_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7; static lean_object* l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg(lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCompletionItem; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6699_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instBEqSymbolKind; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12; static lean_object* l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenModifier_toNat___boxed(lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolTag(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5; static lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3763_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___boxed(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__15; static lean_object* l_Lean_Lsp_instFromJsonCompletionOptions___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolInformation; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_noConfusion(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__15; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9(lean_object*, lean_object*, lean_object*); -static uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4; static lean_object* l_Lean_Lsp_instFromJsonRenameParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4; +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonRenameParams___closed__1; static lean_object* l_Lean_Lsp_instToJsonCompletionList___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8654_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1351____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_SemanticTokenType_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472_(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__9; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionOptions_allCommitCharacters_x3f___default; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3546_(lean_object*); +static uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonFoldingRangeKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__20; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__38; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonRenameOptions; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1(lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlight; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____boxed(lean_object*); lean_object* l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5746_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6270_(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18; LEAN_EXPORT lean_object* l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonCompletionItemKind___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSymbolKind(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278_(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__24; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4; uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339_(uint8_t, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__23; -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDefinitionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentHighlightParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentHighlightParams; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__24; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2; static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__44; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5347_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941_(lean_object*); static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6; lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__6; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23; uint8_t l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_beqRange____x40_Lean_Data_Lsp_Basic___hyg_597_(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3; static lean_object* l_Lean_Lsp_instReprCompletionItemKind___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Lsp_instToJsonSymbolKind___closed__42; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1___boxed(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2; -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770_(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8; +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__3(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5350_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26; static lean_object* l_Lean_Lsp_instBEqSemanticTokenType___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2; lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10; static lean_object* l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1; LEAN_EXPORT lean_object* l_Lean_Lsp_FoldingRangeKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21; static lean_object* l_Lean_Lsp_instToJsonRenameOptions___closed__1; static lean_object* l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; static lean_object* l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____boxed(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96; static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6; static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__10; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10; +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonSemanticTokenType; LEAN_EXPORT lean_object* l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___lambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8651_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonRenameParams; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____boxed(lean_object*); static lean_object* l_Lean_Lsp_SemanticTokenType_names___closed__7; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__5; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6; LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonTypeDefinitionParams; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12; static lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____spec__1___closed__2; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978_(uint8_t); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13; LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonSymbolKind(uint8_t); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244_(lean_object*); +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789_(lean_object*); static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__22; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9; static lean_object* l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7; LEAN_EXPORT lean_object* l_Lean_Lsp_SymbolTag_toCtorIdx___boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47; static lean_object* l_Lean_Lsp_SemanticTokenModifier_names___closed__7; static lean_object* l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12; static lean_object* l_Lean_Lsp_instFromJsonSymbolKind___closed__27; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9; static lean_object* l_Lean_Lsp_instToJsonPrepareRenameParams___closed__1; -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8; -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3(lean_object*, lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107; -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____boxed(lean_object*); -static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22; +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367_(lean_object*); +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; +static lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9; static lean_object* _init_l_Lean_Lsp_CompletionOptions_triggerCharacters_x3f___default() { _start: { @@ -2739,7 +2739,7 @@ x_6 = lean_box(x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__1() { _start: { lean_object* x_1; @@ -2747,17 +2747,17 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.text", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -2766,23 +2766,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2790,7 +2790,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -2799,23 +2799,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2823,7 +2823,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9() { _start: { lean_object* x_1; @@ -2831,33 +2831,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.method", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2865,23 +2865,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2889,7 +2889,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15() { _start: { lean_object* x_1; @@ -2897,33 +2897,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.function", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2931,23 +2931,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2955,7 +2955,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21() { _start: { lean_object* x_1; @@ -2963,33 +2963,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.constructor", 39); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -2997,23 +2997,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3021,7 +3021,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27() { _start: { lean_object* x_1; @@ -3029,33 +3029,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.field", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3063,23 +3063,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3087,7 +3087,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33() { _start: { lean_object* x_1; @@ -3095,33 +3095,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.variable", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3129,23 +3129,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3153,7 +3153,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39() { _start: { lean_object* x_1; @@ -3161,33 +3161,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.class", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3195,23 +3195,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3219,7 +3219,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45() { _start: { lean_object* x_1; @@ -3227,33 +3227,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.interface", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3261,23 +3261,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3285,7 +3285,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51() { _start: { lean_object* x_1; @@ -3293,33 +3293,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.module", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3327,23 +3327,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3351,7 +3351,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57() { _start: { lean_object* x_1; @@ -3359,33 +3359,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.property", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3393,23 +3393,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3417,7 +3417,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63() { _start: { lean_object* x_1; @@ -3425,33 +3425,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.unit", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3459,23 +3459,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3483,7 +3483,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69() { _start: { lean_object* x_1; @@ -3491,33 +3491,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.value", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3525,23 +3525,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3549,7 +3549,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75() { _start: { lean_object* x_1; @@ -3557,33 +3557,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.enum", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3591,23 +3591,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3615,7 +3615,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81() { _start: { lean_object* x_1; @@ -3623,33 +3623,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.keyword", 35); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3657,23 +3657,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3681,7 +3681,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87() { _start: { lean_object* x_1; @@ -3689,33 +3689,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.snippet", 35); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3723,23 +3723,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3747,7 +3747,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93() { _start: { lean_object* x_1; @@ -3755,33 +3755,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.color", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3789,23 +3789,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3813,7 +3813,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99() { _start: { lean_object* x_1; @@ -3821,33 +3821,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.file", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3855,23 +3855,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3879,7 +3879,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105() { _start: { lean_object* x_1; @@ -3887,33 +3887,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.reference", 37); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3921,23 +3921,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3945,7 +3945,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111() { _start: { lean_object* x_1; @@ -3953,33 +3953,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.folder", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -3987,23 +3987,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4011,7 +4011,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117() { _start: { lean_object* x_1; @@ -4019,33 +4019,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.enumMember", 38); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4053,23 +4053,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4077,7 +4077,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123() { _start: { lean_object* x_1; @@ -4085,33 +4085,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.constant", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4119,23 +4119,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4143,7 +4143,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129() { _start: { lean_object* x_1; @@ -4151,33 +4151,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.struct", 34); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4185,23 +4185,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4209,7 +4209,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135() { _start: { lean_object* x_1; @@ -4217,33 +4217,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.event", 33); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4251,23 +4251,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4275,7 +4275,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141() { _start: { lean_object* x_1; @@ -4283,33 +4283,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.operator", 36); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4317,23 +4317,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4341,7 +4341,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147() { _start: { lean_object* x_1; @@ -4349,33 +4349,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Lsp.CompletionItemKind.typeParameter", 41) return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4383,23 +4383,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4407,7 +4407,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -4419,14 +4419,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -4439,14 +4439,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -4459,14 +4459,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20; +x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -4479,14 +4479,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26; +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -4499,14 +4499,14 @@ x_28 = lean_nat_dec_le(x_27, x_2); if (x_28 == 0) { lean_object* x_29; lean_object* x_30; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30; x_30 = l_Repr_addAppParen(x_29, x_2); return x_30; } else { lean_object* x_31; lean_object* x_32; -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32; x_32 = l_Repr_addAppParen(x_31, x_2); return x_32; } @@ -4519,14 +4519,14 @@ x_34 = lean_nat_dec_le(x_33, x_2); if (x_34 == 0) { lean_object* x_35; lean_object* x_36; -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36; x_36 = l_Repr_addAppParen(x_35, x_2); return x_36; } else { lean_object* x_37; lean_object* x_38; -x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38; +x_37 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38; x_38 = l_Repr_addAppParen(x_37, x_2); return x_38; } @@ -4539,14 +4539,14 @@ x_40 = lean_nat_dec_le(x_39, x_2); if (x_40 == 0) { lean_object* x_41; lean_object* x_42; -x_41 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42; +x_41 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42; x_42 = l_Repr_addAppParen(x_41, x_2); return x_42; } else { lean_object* x_43; lean_object* x_44; -x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44; +x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44; x_44 = l_Repr_addAppParen(x_43, x_2); return x_44; } @@ -4559,14 +4559,14 @@ x_46 = lean_nat_dec_le(x_45, x_2); if (x_46 == 0) { lean_object* x_47; lean_object* x_48; -x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48; +x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48; x_48 = l_Repr_addAppParen(x_47, x_2); return x_48; } else { lean_object* x_49; lean_object* x_50; -x_49 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50; +x_49 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50; x_50 = l_Repr_addAppParen(x_49, x_2); return x_50; } @@ -4579,14 +4579,14 @@ x_52 = lean_nat_dec_le(x_51, x_2); if (x_52 == 0) { lean_object* x_53; lean_object* x_54; -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54; x_54 = l_Repr_addAppParen(x_53, x_2); return x_54; } else { lean_object* x_55; lean_object* x_56; -x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56; +x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56; x_56 = l_Repr_addAppParen(x_55, x_2); return x_56; } @@ -4599,14 +4599,14 @@ x_58 = lean_nat_dec_le(x_57, x_2); if (x_58 == 0) { lean_object* x_59; lean_object* x_60; -x_59 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60; +x_59 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60; x_60 = l_Repr_addAppParen(x_59, x_2); return x_60; } else { lean_object* x_61; lean_object* x_62; -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62; x_62 = l_Repr_addAppParen(x_61, x_2); return x_62; } @@ -4619,14 +4619,14 @@ x_64 = lean_nat_dec_le(x_63, x_2); if (x_64 == 0) { lean_object* x_65; lean_object* x_66; -x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66; +x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66; x_66 = l_Repr_addAppParen(x_65, x_2); return x_66; } else { lean_object* x_67; lean_object* x_68; -x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68; +x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68; x_68 = l_Repr_addAppParen(x_67, x_2); return x_68; } @@ -4639,14 +4639,14 @@ x_70 = lean_nat_dec_le(x_69, x_2); if (x_70 == 0) { lean_object* x_71; lean_object* x_72; -x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72; +x_71 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72; x_72 = l_Repr_addAppParen(x_71, x_2); return x_72; } else { lean_object* x_73; lean_object* x_74; -x_73 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74; +x_73 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74; x_74 = l_Repr_addAppParen(x_73, x_2); return x_74; } @@ -4659,14 +4659,14 @@ x_76 = lean_nat_dec_le(x_75, x_2); if (x_76 == 0) { lean_object* x_77; lean_object* x_78; -x_77 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78; +x_77 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78; x_78 = l_Repr_addAppParen(x_77, x_2); return x_78; } else { lean_object* x_79; lean_object* x_80; -x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80; +x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80; x_80 = l_Repr_addAppParen(x_79, x_2); return x_80; } @@ -4679,14 +4679,14 @@ x_82 = lean_nat_dec_le(x_81, x_2); if (x_82 == 0) { lean_object* x_83; lean_object* x_84; -x_83 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84; +x_83 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84; x_84 = l_Repr_addAppParen(x_83, x_2); return x_84; } else { lean_object* x_85; lean_object* x_86; -x_85 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86; +x_85 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86; x_86 = l_Repr_addAppParen(x_85, x_2); return x_86; } @@ -4699,14 +4699,14 @@ x_88 = lean_nat_dec_le(x_87, x_2); if (x_88 == 0) { lean_object* x_89; lean_object* x_90; -x_89 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90; +x_89 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90; x_90 = l_Repr_addAppParen(x_89, x_2); return x_90; } else { lean_object* x_91; lean_object* x_92; -x_91 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92; +x_91 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92; x_92 = l_Repr_addAppParen(x_91, x_2); return x_92; } @@ -4719,14 +4719,14 @@ x_94 = lean_nat_dec_le(x_93, x_2); if (x_94 == 0) { lean_object* x_95; lean_object* x_96; -x_95 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96; +x_95 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96; x_96 = l_Repr_addAppParen(x_95, x_2); return x_96; } else { lean_object* x_97; lean_object* x_98; -x_97 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98; +x_97 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98; x_98 = l_Repr_addAppParen(x_97, x_2); return x_98; } @@ -4739,14 +4739,14 @@ x_100 = lean_nat_dec_le(x_99, x_2); if (x_100 == 0) { lean_object* x_101; lean_object* x_102; -x_101 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102; +x_101 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102; x_102 = l_Repr_addAppParen(x_101, x_2); return x_102; } else { lean_object* x_103; lean_object* x_104; -x_103 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104; +x_103 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104; x_104 = l_Repr_addAppParen(x_103, x_2); return x_104; } @@ -4759,14 +4759,14 @@ x_106 = lean_nat_dec_le(x_105, x_2); if (x_106 == 0) { lean_object* x_107; lean_object* x_108; -x_107 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108; +x_107 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108; x_108 = l_Repr_addAppParen(x_107, x_2); return x_108; } else { lean_object* x_109; lean_object* x_110; -x_109 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110; +x_109 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110; x_110 = l_Repr_addAppParen(x_109, x_2); return x_110; } @@ -4779,14 +4779,14 @@ x_112 = lean_nat_dec_le(x_111, x_2); if (x_112 == 0) { lean_object* x_113; lean_object* x_114; -x_113 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114; +x_113 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114; x_114 = l_Repr_addAppParen(x_113, x_2); return x_114; } else { lean_object* x_115; lean_object* x_116; -x_115 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116; +x_115 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116; x_116 = l_Repr_addAppParen(x_115, x_2); return x_116; } @@ -4799,14 +4799,14 @@ x_118 = lean_nat_dec_le(x_117, x_2); if (x_118 == 0) { lean_object* x_119; lean_object* x_120; -x_119 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120; +x_119 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120; x_120 = l_Repr_addAppParen(x_119, x_2); return x_120; } else { lean_object* x_121; lean_object* x_122; -x_121 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122; +x_121 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122; x_122 = l_Repr_addAppParen(x_121, x_2); return x_122; } @@ -4819,14 +4819,14 @@ x_124 = lean_nat_dec_le(x_123, x_2); if (x_124 == 0) { lean_object* x_125; lean_object* x_126; -x_125 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126; +x_125 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126; x_126 = l_Repr_addAppParen(x_125, x_2); return x_126; } else { lean_object* x_127; lean_object* x_128; -x_127 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128; +x_127 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128; x_128 = l_Repr_addAppParen(x_127, x_2); return x_128; } @@ -4839,14 +4839,14 @@ x_130 = lean_nat_dec_le(x_129, x_2); if (x_130 == 0) { lean_object* x_131; lean_object* x_132; -x_131 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132; +x_131 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132; x_132 = l_Repr_addAppParen(x_131, x_2); return x_132; } else { lean_object* x_133; lean_object* x_134; -x_133 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134; +x_133 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134; x_134 = l_Repr_addAppParen(x_133, x_2); return x_134; } @@ -4859,14 +4859,14 @@ x_136 = lean_nat_dec_le(x_135, x_2); if (x_136 == 0) { lean_object* x_137; lean_object* x_138; -x_137 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138; +x_137 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138; x_138 = l_Repr_addAppParen(x_137, x_2); return x_138; } else { lean_object* x_139; lean_object* x_140; -x_139 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140; +x_139 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140; x_140 = l_Repr_addAppParen(x_139, x_2); return x_140; } @@ -4879,14 +4879,14 @@ x_142 = lean_nat_dec_le(x_141, x_2); if (x_142 == 0) { lean_object* x_143; lean_object* x_144; -x_143 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144; +x_143 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144; x_144 = l_Repr_addAppParen(x_143, x_2); return x_144; } else { lean_object* x_145; lean_object* x_146; -x_145 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146; +x_145 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146; x_146 = l_Repr_addAppParen(x_145, x_2); return x_146; } @@ -4899,14 +4899,14 @@ x_148 = lean_nat_dec_le(x_147, x_2); if (x_148 == 0) { lean_object* x_149; lean_object* x_150; -x_149 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150; +x_149 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150; x_150 = l_Repr_addAppParen(x_149, x_2); return x_150; } else { lean_object* x_151; lean_object* x_152; -x_151 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152; +x_151 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152; x_152 = l_Repr_addAppParen(x_151, x_2); return x_152; } @@ -4914,13 +4914,13 @@ return x_152; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339_(x_3, x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -4929,7 +4929,7 @@ static lean_object* _init_l_Lean_Lsp_instReprCompletionItemKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____boxed), 2, 0); return x_1; } } @@ -5034,7 +5034,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1() { _start: { lean_object* x_1; @@ -5042,7 +5042,7 @@ x_1 = lean_mk_string_from_bytes("newText", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2() { _start: { lean_object* x_1; @@ -5050,78 +5050,78 @@ x_1 = lean_mk_string_from_bytes("InsertReplaceEdit", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10() { _start: { lean_object* x_1; @@ -5129,47 +5129,47 @@ x_1 = lean_mk_string_from_bytes("insert", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15() { _start: { lean_object* x_1; @@ -5177,51 +5177,51 @@ x_1 = lean_mk_string_from_bytes("replace", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -5231,7 +5231,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -5243,7 +5243,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -5257,7 +5257,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -5268,7 +5268,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -5280,7 +5280,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -5294,7 +5294,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15; x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -5306,7 +5306,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -5318,7 +5318,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -5360,11 +5360,11 @@ return x_39; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244_(x_1); lean_dec(x_1); return x_2; } @@ -5373,7 +5373,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____boxed), 1, 0); return x_1; } } @@ -5385,7 +5385,7 @@ x_1 = l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1380_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1383_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -5398,7 +5398,7 @@ lean_inc(x_4); lean_dec(x_1); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_2); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1; x_7 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -5407,7 +5407,7 @@ x_9 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_9, 0, x_7); lean_ctor_set(x_9, 1, x_8); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_3); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -5415,7 +5415,7 @@ x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_8); x_14 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_4); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -5440,7 +5440,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonInsertReplaceEdit___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1380_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1383_), 1, 0); return x_1; } } @@ -5500,7 +5500,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5564,7 +5564,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5640,7 +5640,7 @@ return x_22; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5654,7 +5654,7 @@ return x_4; else { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241_(x_3); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244_(x_3); lean_dec(x_3); if (lean_obj_tag(x_5) == 0) { @@ -5704,7 +5704,7 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -5726,7 +5726,7 @@ return x_6; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1() { _start: { lean_object* x_1; @@ -5734,7 +5734,7 @@ x_1 = lean_mk_string_from_bytes("label", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2() { _start: { lean_object* x_1; @@ -5742,78 +5742,78 @@ x_1 = lean_mk_string_from_bytes("CompletionItem", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10() { _start: { lean_object* x_1; @@ -5821,7 +5821,7 @@ x_1 = lean_mk_string_from_bytes("detail", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11() { _start: { lean_object* x_1; @@ -5829,47 +5829,47 @@ x_1 = lean_mk_string_from_bytes("detail\?", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16() { _start: { lean_object* x_1; @@ -5877,7 +5877,7 @@ x_1 = lean_mk_string_from_bytes("documentation", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17() { _start: { lean_object* x_1; @@ -5885,47 +5885,47 @@ x_1 = lean_mk_string_from_bytes("documentation\?", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22() { _start: { lean_object* x_1; @@ -5933,7 +5933,7 @@ x_1 = lean_mk_string_from_bytes("kind", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23() { _start: { lean_object* x_1; @@ -5941,47 +5941,47 @@ x_1 = lean_mk_string_from_bytes("kind\?", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28() { _start: { lean_object* x_1; @@ -5989,7 +5989,7 @@ x_1 = lean_mk_string_from_bytes("textEdit", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29() { _start: { lean_object* x_1; @@ -5997,47 +5997,47 @@ x_1 = lean_mk_string_from_bytes("textEdit\?", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34() { _start: { lean_object* x_1; @@ -6045,7 +6045,7 @@ x_1 = lean_mk_string_from_bytes("sortText", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35() { _start: { lean_object* x_1; @@ -6053,47 +6053,47 @@ x_1 = lean_mk_string_from_bytes("sortText\?", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40() { _start: { lean_object* x_1; @@ -6101,7 +6101,7 @@ x_1 = lean_mk_string_from_bytes("data", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41() { _start: { lean_object* x_1; @@ -6109,51 +6109,51 @@ x_1 = lean_mk_string_from_bytes("data\?", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__45() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6163,7 +6163,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6175,7 +6175,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6189,7 +6189,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -6200,7 +6200,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -6212,7 +6212,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -6226,8 +6226,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -6238,7 +6238,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -6250,7 +6250,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -6264,8 +6264,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -6277,7 +6277,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -6289,7 +6289,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -6303,8 +6303,8 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28; -x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3(x_1, x_46); +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28; +x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3(x_1, x_46); if (lean_obj_tag(x_47) == 0) { uint8_t x_48; @@ -6317,7 +6317,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -6329,7 +6329,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -6343,7 +6343,7 @@ lean_object* x_56; lean_object* x_57; lean_object* x_58; x_56 = lean_ctor_get(x_47, 0); lean_inc(x_56); lean_dec(x_47); -x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34; +x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34; x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { @@ -6358,7 +6358,7 @@ if (x_59 == 0) { lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_58, 0); -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39; x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); lean_ctor_set(x_58, 0, x_62); @@ -6370,7 +6370,7 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; x_63 = lean_ctor_get(x_58, 0); lean_inc(x_63); lean_dec(x_58); -x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39; +x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39; x_65 = lean_string_append(x_64, x_63); lean_dec(x_63); x_66 = lean_alloc_ctor(0, 1, 0); @@ -6384,8 +6384,8 @@ lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; x_67 = lean_ctor_get(x_58, 0); lean_inc(x_67); lean_dec(x_58); -x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; -x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4(x_1, x_68); +x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; +x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4(x_1, x_68); x_70 = !lean_is_exclusive(x_69); if (x_70 == 0) { @@ -6428,51 +6428,51 @@ return x_75; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__2(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__2(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__3(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514_(x_1); lean_dec(x_1); return x_2; } @@ -6481,7 +6481,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____boxed), 1, 0); return x_1; } } @@ -6493,7 +6493,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6519,7 +6519,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6555,7 +6555,7 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -6571,7 +6571,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_obj x_4 = lean_ctor_get(x_2, 0); lean_inc(x_4); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1380_(x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1383_(x_4); x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_1); lean_ctor_set(x_6, 1, x_5); @@ -6583,7 +6583,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -6604,7 +6604,7 @@ lean_inc(x_8); lean_dec(x_1); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_2); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -6612,20 +6612,20 @@ x_12 = lean_box(0); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_14, x_3); lean_dec(x_3); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; -x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1(x_16, x_4); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; +x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1(x_16, x_4); lean_dec(x_4); -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__2(x_18, x_5); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28; -x_21 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__3(x_20, x_6); -x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34; +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_19 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__2(x_18, x_5); +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28; +x_21 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__3(x_20, x_6); +x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34; x_23 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_22, x_7); lean_dec(x_7); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; x_25 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_24, x_8); lean_dec(x_8); x_26 = lean_alloc_ctor(1, 2, 0); @@ -6654,11 +6654,11 @@ x_34 = l_Lean_Json_mkObj(x_33); return x_34; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -6667,7 +6667,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809_), 1, 0); return x_1; } } @@ -6712,7 +6712,7 @@ x_1 = l_Lean_Lsp_instInhabitedCompletionItem___closed__2; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -6730,7 +6730,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_6 = lean_array_uget(x_3, x_2); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_3, x_2, x_7); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511_(x_6); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514_(x_6); lean_dec(x_6); if (lean_obj_tag(x_9) == 0) { @@ -6768,7 +6768,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -6785,7 +6785,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__2(x_6, x_7, x_4); return x_8; } else @@ -6804,7 +6804,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1() { _start: { lean_object* x_1; @@ -6812,7 +6812,7 @@ x_1 = lean_mk_string_from_bytes("isIncomplete", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2() { _start: { lean_object* x_1; @@ -6820,78 +6820,78 @@ x_1 = lean_mk_string_from_bytes("CompletionList", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10() { _start: { lean_object* x_1; @@ -6899,51 +6899,51 @@ x_1 = lean_mk_string_from_bytes("items", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -6954,7 +6954,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -6966,7 +6966,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -6980,8 +6980,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -6991,7 +6991,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7003,7 +7003,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7046,7 +7046,7 @@ return x_30; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____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; @@ -7054,7 +7054,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____spec__2(x_4, x_5, x_3); return x_6; } } @@ -7062,7 +7062,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918_), 1, 0); return x_1; } } @@ -7074,7 +7074,7 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionList___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -7089,7 +7089,7 @@ lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; size_t x x_5 = lean_array_uget(x_3, x_2); x_6 = lean_unsigned_to_nat(0u); x_7 = lean_array_uset(x_3, x_2, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1806_(x_5); +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1809_(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); @@ -7099,7 +7099,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -7109,7 +7109,7 @@ lean_inc(x_3); lean_dec(x_1); x_4 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_4, 0, x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -7121,10 +7121,10 @@ x_9 = lean_array_get_size(x_3); x_10 = lean_usize_of_nat(x_9); lean_dec(x_9); x_11 = 0; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1(x_10, x_11, x_3); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____spec__1(x_10, x_11, x_3); x_13 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -7142,7 +7142,7 @@ x_20 = l_Lean_Json_mkObj(x_19); return x_20; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____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; @@ -7150,7 +7150,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018____spec__1(x_4, x_5, x_3); return x_6; } } @@ -7158,7 +7158,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionList___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2015_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2018_), 1, 0); return x_1; } } @@ -7170,7 +7170,7 @@ x_1 = l_Lean_Lsp_instToJsonCompletionList___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1() { _start: { lean_object* x_1; @@ -7178,7 +7178,7 @@ x_1 = lean_mk_string_from_bytes("textDocument", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2() { _start: { lean_object* x_1; @@ -7186,78 +7186,78 @@ x_1 = lean_mk_string_from_bytes("CompletionParams", 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10() { _start: { lean_object* x_1; @@ -7265,51 +7265,51 @@ x_1 = lean_mk_string_from_bytes("position", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -7319,7 +7319,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7331,7 +7331,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7345,7 +7345,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -7356,7 +7356,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7368,7 +7368,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7407,11 +7407,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075_(x_1); lean_dec(x_1); return x_2; } @@ -7420,7 +7420,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____boxed), 1, 0); return x_1; } } @@ -7432,14 +7432,14 @@ x_1 = l_Lean_Lsp_instFromJsonCompletionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2175_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -7451,7 +7451,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -7473,7 +7473,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCompletionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2172_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2175_), 1, 0); return x_1; } } @@ -7493,7 +7493,7 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1() { _start: { lean_object* x_1; @@ -7501,7 +7501,7 @@ x_1 = lean_mk_string_from_bytes("contents", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2() { _start: { lean_object* x_1; @@ -7509,7 +7509,7 @@ x_1 = lean_mk_string_from_bytes("range", 5); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -7517,7 +7517,7 @@ x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonMarkupContent____x40_Lean_Data_Lsp_Basic___hyg_5746_(x_2); lean_dec(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -7528,7 +7528,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1284____spec__1(x_9, x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -7545,7 +7545,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242_), 1, 0); return x_1; } } @@ -7557,7 +7557,7 @@ x_1 = l_Lean_Lsp_instToJsonHover___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -7567,7 +7567,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1() { _start: { lean_object* x_1; @@ -7575,78 +7575,78 @@ x_1 = lean_mk_string_from_bytes("Hover", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9() { _start: { lean_object* x_1; @@ -7654,52 +7654,52 @@ x_1 = lean_mk_string_from_bytes("range\?", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -7708,7 +7708,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7720,7 +7720,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7734,7 +7734,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocationLink____x40_Lean_Data_Lsp_Basic___hyg_1351____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -7745,7 +7745,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7757,7 +7757,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -7796,21 +7796,21 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281_(x_1); lean_dec(x_1); return x_2; } @@ -7819,7 +7819,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHover___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____boxed), 1, 0); return x_1; } } @@ -7831,7 +7831,7 @@ x_1 = l_Lean_Lsp_instFromJsonHover___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1() { _start: { lean_object* x_1; @@ -7839,82 +7839,82 @@ x_1 = lean_mk_string_from_bytes("HoverParams", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -7924,7 +7924,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -7936,7 +7936,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -7950,7 +7950,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -7961,7 +7961,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -7973,7 +7973,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8012,11 +8012,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392_(x_1); lean_dec(x_1); return x_2; } @@ -8025,7 +8025,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____boxed), 1, 0); return x_1; } } @@ -8037,14 +8037,14 @@ x_1 = l_Lean_Lsp_instFromJsonHoverParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2489_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2492_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8056,7 +8056,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -8078,7 +8078,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonHoverParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2489_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2492_), 1, 0); return x_1; } } @@ -8090,7 +8090,7 @@ x_1 = l_Lean_Lsp_instToJsonHoverParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1() { _start: { lean_object* x_1; @@ -8098,82 +8098,82 @@ x_1 = lean_mk_string_from_bytes("DeclarationParams", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -8183,7 +8183,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8195,7 +8195,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8209,7 +8209,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -8220,7 +8220,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8232,7 +8232,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8271,11 +8271,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549_(x_1); lean_dec(x_1); return x_2; } @@ -8284,7 +8284,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDeclarationParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____boxed), 1, 0); return x_1; } } @@ -8296,14 +8296,14 @@ x_1 = l_Lean_Lsp_instFromJsonDeclarationParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2646_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2649_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8315,7 +8315,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -8337,7 +8337,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDeclarationParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2646_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2649_), 1, 0); return x_1; } } @@ -8349,7 +8349,7 @@ x_1 = l_Lean_Lsp_instToJsonDeclarationParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1() { _start: { lean_object* x_1; @@ -8357,82 +8357,82 @@ x_1 = lean_mk_string_from_bytes("DefinitionParams", 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -8442,7 +8442,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8454,7 +8454,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8468,7 +8468,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -8479,7 +8479,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8491,7 +8491,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8530,11 +8530,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706_(x_1); lean_dec(x_1); return x_2; } @@ -8543,7 +8543,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDefinitionParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____boxed), 1, 0); return x_1; } } @@ -8555,14 +8555,14 @@ x_1 = l_Lean_Lsp_instFromJsonDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2803_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2806_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8574,7 +8574,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -8596,7 +8596,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDefinitionParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2803_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2806_), 1, 0); return x_1; } } @@ -8608,7 +8608,7 @@ x_1 = l_Lean_Lsp_instToJsonDefinitionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1() { _start: { lean_object* x_1; @@ -8616,82 +8616,82 @@ x_1 = lean_mk_string_from_bytes("TypeDefinitionParams", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -8701,7 +8701,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8713,7 +8713,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -8727,7 +8727,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -8738,7 +8738,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -8750,7 +8750,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -8789,11 +8789,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863_(x_1); lean_dec(x_1); return x_2; } @@ -8802,7 +8802,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____boxed), 1, 0); return x_1; } } @@ -8814,14 +8814,14 @@ x_1 = l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2960_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2963_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -8833,7 +8833,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -8855,7 +8855,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2960_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2963_), 1, 0); return x_1; } } @@ -8867,7 +8867,7 @@ x_1 = l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1() { _start: { lean_object* x_1; @@ -8875,7 +8875,7 @@ x_1 = lean_mk_string_from_bytes("includeDeclaration", 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2() { _start: { lean_object* x_1; @@ -8883,82 +8883,82 @@ x_1 = lean_mk_string_from_bytes("ReferenceContext", 16); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -8968,7 +8968,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -8980,7 +8980,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9009,11 +9009,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020_(x_1); lean_dec(x_1); return x_2; } @@ -9022,7 +9022,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____boxed), 1, 0); return x_1; } } @@ -9034,13 +9034,13 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -9056,13 +9056,13 @@ x_9 = l_Lean_Json_mkObj(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081_(x_2); return x_3; } } @@ -9070,7 +9070,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceContext___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081____boxed), 1, 0); return x_1; } } @@ -9082,17 +9082,17 @@ x_1 = l_Lean_Lsp_instToJsonReferenceContext___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020_(x_3); lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1() { _start: { lean_object* x_1; @@ -9100,78 +9100,78 @@ x_1 = lean_mk_string_from_bytes("ReferenceParams", 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9() { _start: { lean_object* x_1; @@ -9179,51 +9179,51 @@ x_1 = lean_mk_string_from_bytes("context", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -9233,7 +9233,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9245,7 +9245,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9259,7 +9259,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9270,7 +9270,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9282,7 +9282,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9296,8 +9296,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -9308,7 +9308,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -9320,7 +9320,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -9366,21 +9366,21 @@ return x_41; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132_(x_1); lean_dec(x_1); return x_2; } @@ -9389,7 +9389,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____boxed), 1, 0); return x_1; } } @@ -9401,7 +9401,7 @@ x_1 = l_Lean_Lsp_instFromJsonReferenceParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3268_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3271_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -9410,7 +9410,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -9422,7 +9422,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -9434,8 +9434,8 @@ lean_inc(x_14); lean_dec(x_1); x_15 = lean_unbox(x_14); lean_dec(x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3078_(x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3081_(x_15); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -9460,7 +9460,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonReferenceParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3268_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3271_), 1, 0); return x_1; } } @@ -9472,7 +9472,7 @@ x_1 = l_Lean_Lsp_instToJsonReferenceParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1() { _start: { lean_object* x_1; @@ -9480,7 +9480,7 @@ x_1 = lean_mk_string_from_bytes("query", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2() { _start: { lean_object* x_1; @@ -9488,82 +9488,82 @@ x_1 = lean_mk_string_from_bytes("WorkspaceSymbolParams", 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -9573,7 +9573,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9585,7 +9585,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9614,11 +9614,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342_(x_1); lean_dec(x_1); return x_2; } @@ -9627,7 +9627,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____boxed), 1, 0); return x_1; } } @@ -9639,13 +9639,13 @@ x_1 = l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3400_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3403_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -9665,7 +9665,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3400_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3403_), 1, 0); return x_1; } } @@ -9677,7 +9677,7 @@ x_1 = l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1() { _start: { lean_object* x_1; @@ -9685,82 +9685,82 @@ x_1 = lean_mk_string_from_bytes("DocumentHighlightParams", 23); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -9770,7 +9770,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -9782,7 +9782,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -9796,7 +9796,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -9807,7 +9807,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -9819,7 +9819,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -9858,11 +9858,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446_(x_1); lean_dec(x_1); return x_2; } @@ -9871,7 +9871,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____boxed), 1, 0); return x_1; } } @@ -9883,14 +9883,14 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3543_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3546_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -9902,7 +9902,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -9924,7 +9924,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlightParams___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3543_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3546_), 1, 0); return x_1; } } @@ -10099,7 +10099,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -10156,14 +10156,14 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -10174,8 +10174,8 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1(x_9, x_8); +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_10 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_10); @@ -10188,11 +10188,11 @@ x_14 = l_Lean_Json_mkObj(x_13); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -10201,7 +10201,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentHighlight___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3645_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentHighlight____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3648_), 1, 0); return x_1; } } @@ -10213,7 +10213,7 @@ x_1 = l_Lean_Lsp_instToJsonDocumentHighlight___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1() { _start: { lean_object* x_1; @@ -10221,62 +10221,62 @@ x_1 = lean_mk_string_from_bytes("DocumentSymbolParams", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -10286,7 +10286,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -10298,7 +10298,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -10327,11 +10327,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702_(x_1); lean_dec(x_1); return x_2; } @@ -10340,7 +10340,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____boxed), 1, 0); return x_1; } } @@ -10352,12 +10352,12 @@ x_1 = l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3760_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3763_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -10377,7 +10377,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonDocumentSymbolParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3760_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3763_), 1, 0); return x_1; } } @@ -10590,7 +10590,7 @@ x_6 = l_Lean_Lsp_SymbolKind_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -10602,7 +10602,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799____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; @@ -10610,7 +10610,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796_(x_3, x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -10619,7 +10619,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSymbolKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799____boxed), 2, 0); return x_1; } } @@ -10631,7 +10631,7 @@ x_1 = l_Lean_Lsp_instBEqSymbolKind___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815_(uint8_t x_1) { _start: { switch (x_1) { @@ -10794,13 +10794,13 @@ return x_27; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815____boxed(lean_object* x_1) { _start: { uint8_t x_2; uint64_t x_3; lean_object* x_4; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -10809,7 +10809,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSymbolKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815____boxed), 1, 0); return x_1; } } @@ -13279,7 +13279,7 @@ x_671 = lean_nat_dec_eq(x_5, x_670); if (x_671 == 0) { lean_object* x_672; lean_object* x_673; lean_object* x_674; lean_object* x_675; lean_object* x_676; lean_object* x_677; lean_object* x_678; lean_object* x_679; -x_672 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_672 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; lean_ctor_set(x_2, 0, x_672); x_673 = lean_unsigned_to_nat(80u); x_674 = l_Lean_Json_pretty(x_1, x_673); @@ -13311,7 +13311,7 @@ x_682 = lean_nat_dec_eq(x_5, x_681); if (x_682 == 0) { lean_object* x_683; lean_object* x_684; lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; -x_683 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_683 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; lean_ctor_set(x_2, 0, x_683); x_684 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_684, 0, x_2); @@ -13352,7 +13352,7 @@ x_696 = lean_nat_dec_eq(x_5, x_695); if (x_696 == 0) { lean_object* x_697; lean_object* x_698; lean_object* x_699; lean_object* x_700; lean_object* x_701; lean_object* x_702; lean_object* x_703; lean_object* x_704; -x_697 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_697 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_2, 0, x_697); x_698 = lean_unsigned_to_nat(80u); x_699 = l_Lean_Json_pretty(x_1, x_698); @@ -13384,7 +13384,7 @@ x_707 = lean_nat_dec_eq(x_5, x_706); if (x_707 == 0) { lean_object* x_708; lean_object* x_709; lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; -x_708 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_708 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_2, 0, x_708); x_709 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_709, 0, x_2); @@ -14729,7 +14729,7 @@ x_1127 = lean_nat_dec_eq(x_726, x_1126); if (x_1127 == 0) { lean_object* x_1128; lean_object* x_1129; lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; -x_1128 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_1128 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; x_1129 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1129, 0, x_1128); lean_ctor_set(x_1129, 1, x_726); @@ -14776,7 +14776,7 @@ x_1141 = lean_nat_dec_eq(x_726, x_1140); if (x_1141 == 0) { lean_object* x_1142; lean_object* x_1143; lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; -x_1142 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_1142 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; x_1143 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1143, 0, x_1142); lean_ctor_set(x_1143, 1, x_726); @@ -15467,7 +15467,7 @@ x_2 = lean_box(0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -17400,7 +17400,7 @@ x_673 = lean_nat_dec_eq(x_7, x_672); if (x_673 == 0) { 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; -x_674 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_674 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; lean_ctor_set(x_4, 0, x_674); x_675 = lean_unsigned_to_nat(80u); x_676 = l_Lean_Json_pretty(x_3, x_675); @@ -17432,7 +17432,7 @@ x_684 = lean_nat_dec_eq(x_7, x_683); if (x_684 == 0) { lean_object* x_685; lean_object* x_686; lean_object* x_687; lean_object* x_688; lean_object* x_689; lean_object* x_690; lean_object* x_691; lean_object* x_692; lean_object* x_693; -x_685 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_685 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; lean_ctor_set(x_4, 0, x_685); x_686 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_686, 0, x_4); @@ -17473,7 +17473,7 @@ x_698 = lean_nat_dec_eq(x_7, x_697); if (x_698 == 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_object* x_706; -x_699 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_699 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_4, 0, x_699); x_700 = lean_unsigned_to_nat(80u); x_701 = l_Lean_Json_pretty(x_3, x_700); @@ -17505,7 +17505,7 @@ x_709 = lean_nat_dec_eq(x_7, x_708); if (x_709 == 0) { lean_object* x_710; lean_object* x_711; lean_object* x_712; lean_object* x_713; lean_object* x_714; lean_object* x_715; lean_object* x_716; lean_object* x_717; lean_object* x_718; -x_710 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_710 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_4, 0, x_710); x_711 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_711, 0, x_4); @@ -18850,7 +18850,7 @@ x_1129 = lean_nat_dec_eq(x_728, x_1128); if (x_1129 == 0) { lean_object* x_1130; lean_object* x_1131; lean_object* x_1132; lean_object* x_1133; lean_object* x_1134; lean_object* x_1135; lean_object* x_1136; lean_object* x_1137; lean_object* x_1138; lean_object* x_1139; -x_1130 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3; +x_1130 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3; x_1131 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1131, 0, x_1130); lean_ctor_set(x_1131, 1, x_728); @@ -18897,7 +18897,7 @@ x_1143 = lean_nat_dec_eq(x_728, x_1142); if (x_1143 == 0) { lean_object* x_1144; lean_object* x_1145; lean_object* x_1146; lean_object* x_1147; lean_object* x_1148; lean_object* x_1149; lean_object* x_1150; lean_object* x_1151; lean_object* x_1152; lean_object* x_1153; -x_1144 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_1144 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; x_1145 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_1145, 0, x_1144); lean_ctor_set(x_1145, 1, x_728); @@ -18962,7 +18962,7 @@ return x_1168; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -19020,15 +19020,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; @@ -19051,7 +19051,7 @@ x_7 = lean_array_get_size(x_6); x_8 = lean_usize_of_nat(x_7); lean_dec(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg(x_1, x_8, x_9, x_6); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg(x_1, x_8, x_9, x_6); if (lean_obj_tag(x_10) == 0) { uint8_t x_11; @@ -19116,15 +19116,15 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg___boxed), 3, 0); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1() { _start: { lean_object* x_1; @@ -19132,7 +19132,7 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2() { _start: { lean_object* x_1; @@ -19140,178 +19140,178 @@ x_1 = lean_mk_string_from_bytes("DocumentSymbolAux", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20() { _start: { lean_object* x_1; @@ -19319,47 +19319,47 @@ x_1 = lean_mk_string_from_bytes("selectionRange", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25() { _start: { lean_object* x_1; @@ -19367,7 +19367,7 @@ x_1 = lean_mk_string_from_bytes("children", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26() { _start: { lean_object* x_1; @@ -19375,51 +19375,51 @@ x_1 = lean_mk_string_from_bytes("children\?", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_4 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_2, x_3); if (lean_obj_tag(x_4) == 0) { @@ -19430,7 +19430,7 @@ if (x_5 == 0) { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_4, 0); -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9; x_8 = lean_string_append(x_7, x_6); lean_dec(x_6); lean_ctor_set(x_4, 0, x_8); @@ -19442,7 +19442,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_9 = lean_ctor_get(x_4, 0); lean_inc(x_9); lean_dec(x_4); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9; x_11 = lean_string_append(x_10, x_9); lean_dec(x_9); x_12 = lean_alloc_ctor(0, 1, 0); @@ -19456,7 +19456,7 @@ lean_object* x_13; lean_object* x_14; lean_object* x_15; x_13 = lean_ctor_get(x_4, 0); lean_inc(x_13); lean_dec(x_4); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_15 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_2, x_14); if (lean_obj_tag(x_15) == 0) { @@ -19468,7 +19468,7 @@ if (x_16 == 0) { lean_object* x_17; lean_object* x_18; lean_object* x_19; x_17 = lean_ctor_get(x_15, 0); -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11; +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11; x_19 = lean_string_append(x_18, x_17); lean_dec(x_17); lean_ctor_set(x_15, 0, x_19); @@ -19480,7 +19480,7 @@ lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; x_20 = lean_ctor_get(x_15, 0); lean_inc(x_20); lean_dec(x_15); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11; x_22 = lean_string_append(x_21, x_20); lean_dec(x_20); x_23 = lean_alloc_ctor(0, 1, 0); @@ -19494,8 +19494,8 @@ lean_object* x_24; lean_object* x_25; lean_object* x_26; x_24 = lean_ctor_get(x_15, 0); lean_inc(x_24); lean_dec(x_15); -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(x_2, x_25); +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_26 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1(x_2, x_25); if (lean_obj_tag(x_26) == 0) { uint8_t x_27; @@ -19507,7 +19507,7 @@ if (x_27 == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; x_28 = lean_ctor_get(x_26, 0); -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15; x_30 = lean_string_append(x_29, x_28); lean_dec(x_28); lean_ctor_set(x_26, 0, x_30); @@ -19519,7 +19519,7 @@ lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; x_31 = lean_ctor_get(x_26, 0); lean_inc(x_31); lean_dec(x_26); -x_32 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15; +x_32 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15; x_33 = lean_string_append(x_32, x_31); lean_dec(x_31); x_34 = lean_alloc_ctor(0, 1, 0); @@ -19533,7 +19533,7 @@ lean_object* x_35; lean_object* x_36; lean_object* x_37; x_35 = lean_ctor_get(x_26, 0); lean_inc(x_35); lean_dec(x_26); -x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_36 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_37 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_2, x_36); if (lean_obj_tag(x_37) == 0) { @@ -19547,7 +19547,7 @@ if (x_38 == 0) { lean_object* x_39; lean_object* x_40; lean_object* x_41; x_39 = lean_ctor_get(x_37, 0); -x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19; +x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19; x_41 = lean_string_append(x_40, x_39); lean_dec(x_39); lean_ctor_set(x_37, 0, x_41); @@ -19559,7 +19559,7 @@ lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; x_42 = lean_ctor_get(x_37, 0); lean_inc(x_42); lean_dec(x_37); -x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19; +x_43 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19; x_44 = lean_string_append(x_43, x_42); lean_dec(x_42); x_45 = lean_alloc_ctor(0, 1, 0); @@ -19573,7 +19573,7 @@ lean_object* x_46; lean_object* x_47; lean_object* x_48; x_46 = lean_ctor_get(x_37, 0); lean_inc(x_46); lean_dec(x_37); -x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_47 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_48 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_2, x_47); if (lean_obj_tag(x_48) == 0) { @@ -19588,7 +19588,7 @@ if (x_49 == 0) { lean_object* x_50; lean_object* x_51; lean_object* x_52; x_50 = lean_ctor_get(x_48, 0); -x_51 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24; +x_51 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24; x_52 = lean_string_append(x_51, x_50); lean_dec(x_50); lean_ctor_set(x_48, 0, x_52); @@ -19600,7 +19600,7 @@ lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; x_53 = lean_ctor_get(x_48, 0); lean_inc(x_53); lean_dec(x_48); -x_54 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24; +x_54 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24; x_55 = lean_string_append(x_54, x_53); lean_dec(x_53); x_56 = lean_alloc_ctor(0, 1, 0); @@ -19614,8 +19614,8 @@ lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_48, 0); lean_inc(x_57); lean_dec(x_48); -x_58 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25; -x_59 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg(x_1, x_2, x_58); +x_58 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25; +x_59 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg(x_1, x_2, x_58); if (lean_obj_tag(x_59) == 0) { uint8_t x_60; @@ -19629,7 +19629,7 @@ if (x_60 == 0) { lean_object* x_61; lean_object* x_62; lean_object* x_63; x_61 = lean_ctor_get(x_59, 0); -x_62 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30; +x_62 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30; x_63 = lean_string_append(x_62, x_61); lean_dec(x_61); lean_ctor_set(x_59, 0, x_63); @@ -19641,7 +19641,7 @@ lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; x_64 = lean_ctor_get(x_59, 0); lean_inc(x_64); lean_dec(x_59); -x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30; +x_65 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30; x_66 = lean_string_append(x_65, x_64); lean_dec(x_64); x_67 = lean_alloc_ctor(0, 1, 0); @@ -19696,25 +19696,25 @@ return x_75; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___boxed), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___boxed), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -19722,25 +19722,25 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__3___rarg(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__3___rarg(x_1, x_5, x_6, x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__2___rarg(x_1, x_2, x_3); +x_4 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__2___rarg(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg(x_1, x_2); lean_dec(x_2); return x_3; } @@ -19749,7 +19749,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg(lean_ob _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___boxed), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___boxed), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -19762,7 +19762,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Lsp_instFromJsonDocumentSymbolAux___rarg return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -19789,15 +19789,15 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { if (lean_obj_tag(x_3) == 0) @@ -19818,7 +19818,7 @@ x_6 = lean_array_get_size(x_5); x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg(x_1, x_7, x_8, x_5); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg(x_1, x_7, x_8, x_5); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); x_11 = lean_alloc_ctor(0, 2, 0); @@ -19832,15 +19832,15 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____rarg(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____rarg(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -19858,7 +19858,7 @@ lean_inc(x_8); lean_dec(x_2); x_9 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_9, 0, x_3); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -19866,11 +19866,11 @@ x_12 = lean_box(0); x_13 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_15 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_14, x_4); lean_dec(x_4); x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_6); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -19878,15 +19878,15 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_12); x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_7); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_12); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25; -x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__1___rarg(x_1, x_24, x_8); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25; +x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__1___rarg(x_1, x_24, x_8); x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_12); @@ -20083,7 +20083,7 @@ goto block_38; block_38: { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -20105,15 +20105,15 @@ return x_37; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595_(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____rarg), 2, 0); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____rarg), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { size_t x_5; size_t x_6; lean_object* x_7; @@ -20121,7 +20121,7 @@ x_5 = lean_unbox_usize(x_2); lean_dec(x_2); x_6 = lean_unbox_usize(x_3); lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____spec__2___rarg(x_1, x_5, x_6, x_4); +x_7 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____spec__2___rarg(x_1, x_5, x_6, x_4); return x_7; } } @@ -20129,7 +20129,7 @@ LEAN_EXPORT lean_object* l_Lean_Lsp_instToJsonDocumentSymbolAux___rarg(lean_obje _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____rarg), 2, 1); +x_2 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____rarg), 2, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -20201,7 +20201,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; @@ -20219,7 +20219,7 @@ lean_inc(x_7); lean_dec(x_1); x_8 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_8, 0, x_2); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -20227,11 +20227,11 @@ x_11 = lean_box(0); x_12 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_12, 0, x_10); lean_ctor_set(x_12, 1, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_14 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_13, x_3); lean_dec(x_3); x_15 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_5); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -20239,14 +20239,14 @@ x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_11); x_19 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_6); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_11); -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25; x_24 = l_Lean_Json_opt___at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__2(x_23, x_7); x_25 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_25, 0, x_24); @@ -20444,7 +20444,7 @@ goto block_37; block_37: { lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +x_29 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_28); @@ -20473,7 +20473,7 @@ lean_object* x_2; lean_object* x_3; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4592____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4595____at_Lean_Lsp_instToJsonDocumentSymbol_go___spec__1(x_2); return x_3; } } @@ -20570,7 +20570,7 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773_(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -20578,11 +20578,11 @@ x_3 = 1; return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -20593,7 +20593,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSymbolTag___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4770____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4773____boxed), 2, 0); return x_1; } } @@ -20605,7 +20605,7 @@ x_1 = l_Lean_Lsp_instBEqSymbolTag___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789_(lean_object* x_1) { _start: { uint64_t x_2; @@ -20613,11 +20613,11 @@ x_2 = 0; return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -20627,7 +20627,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSymbolTag___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4786____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolTag____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4789____boxed), 1, 0); return x_1; } } @@ -20719,7 +20719,7 @@ x_21 = lean_nat_dec_eq(x_5, x_20); if (x_21 == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_2, 0, x_22); x_23 = lean_unsigned_to_nat(80u); x_24 = l_Lean_Json_pretty(x_1, x_23); @@ -20751,7 +20751,7 @@ x_32 = lean_nat_dec_eq(x_5, x_31); if (x_32 == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_33 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_33 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_2, 0, x_33); x_34 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_34, 0, x_2); @@ -20843,7 +20843,7 @@ x_66 = lean_nat_dec_eq(x_51, x_65); if (x_66 == 0) { lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_67 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; x_68 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_68, 0, x_67); lean_ctor_set(x_68, 1, x_51); @@ -20950,7 +20950,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -21020,7 +21020,7 @@ 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_dec(x_9); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_7, 0, x_28); x_29 = lean_unsigned_to_nat(80u); x_30 = l_Lean_Json_pretty(x_6, x_29); @@ -21057,7 +21057,7 @@ if (x_41 == 0) { lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_dec(x_9); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; lean_ctor_set(x_7, 0, x_42); x_43 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_43, 0, x_7); @@ -21156,7 +21156,7 @@ if (x_78 == 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; lean_object* x_88; lean_dec(x_9); -x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6; +x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6; x_80 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_80, 0, x_79); lean_ctor_set(x_80, 1, x_64); @@ -21229,7 +21229,7 @@ return x_107; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -21246,7 +21246,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2(x_6, x_7, x_4); return x_8; } else @@ -21265,7 +21265,7 @@ return x_15; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -21275,7 +21275,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1() { _start: { lean_object* x_1; @@ -21283,78 +21283,78 @@ x_1 = lean_mk_string_from_bytes("SymbolInformation", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9() { _start: { lean_object* x_1; @@ -21362,47 +21362,47 @@ x_1 = lean_mk_string_from_bytes("tags", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14() { _start: { lean_object* x_1; @@ -21410,47 +21410,47 @@ x_1 = lean_mk_string_from_bytes("location", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19() { _start: { lean_object* x_1; @@ -21458,7 +21458,7 @@ x_1 = lean_mk_string_from_bytes("containerName", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20() { _start: { lean_object* x_1; @@ -21466,51 +21466,51 @@ x_1 = lean_mk_string_from_bytes("containerName\?", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -21521,7 +21521,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -21533,7 +21533,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -21547,8 +21547,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -21559,7 +21559,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -21571,7 +21571,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -21585,9 +21585,9 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; lean_inc(x_1); -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__1(x_1, x_24); +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -21599,7 +21599,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -21611,7 +21611,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -21625,8 +21625,8 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14; -x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3(x_1, x_35); +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14; +x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3(x_1, x_35); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -21639,7 +21639,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -21651,7 +21651,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -21665,7 +21665,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19; +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19; x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_1, x_46); lean_dec(x_1); if (lean_obj_tag(x_47) == 0) @@ -21680,7 +21680,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -21692,7 +21692,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -21744,7 +21744,7 @@ return x_63; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____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; @@ -21752,15 +21752,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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__3(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__3(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -21770,7 +21770,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSymbolInformation___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941_), 1, 0); return x_1; } } @@ -21782,7 +21782,7 @@ x_1 = l_Lean_Lsp_instFromJsonSymbolInformation___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -21806,7 +21806,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; @@ -21822,7 +21822,7 @@ lean_inc(x_6); lean_dec(x_1); x_7 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_7, 0, x_2); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -21834,10 +21834,10 @@ x_12 = lean_array_get_size(x_4); x_13 = lean_usize_of_nat(x_12); lean_dec(x_12); x_14 = 0; -x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1(x_13, x_14, x_4); +x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1(x_13, x_14, x_4); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -21845,14 +21845,14 @@ x_19 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_10); x_20 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1038_(x_5); -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14; x_22 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_20); x_23 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_10); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19; x_25 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_24, x_6); lean_dec(x_6); x_26 = lean_alloc_ctor(1, 2, 0); @@ -22051,7 +22051,7 @@ goto block_37; block_37: { lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +x_30 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -22070,7 +22070,7 @@ return x_36; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____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; @@ -22078,7 +22078,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1(x_4, x_5, x_3); return x_6; } } @@ -22086,7 +22086,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSymbolInformation___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158_), 1, 0); return x_1; } } @@ -22098,7 +22098,7 @@ x_1 = l_Lean_Lsp_instToJsonSymbolInformation___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1() { _start: { lean_object* x_1; @@ -22106,82 +22106,82 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyPrepareParams", 26); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -22191,7 +22191,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22203,7 +22203,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22217,7 +22217,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -22228,7 +22228,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -22240,7 +22240,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -22279,11 +22279,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250_(x_1); lean_dec(x_1); return x_2; } @@ -22292,7 +22292,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___cl _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____boxed), 1, 0); return x_1; } } @@ -22304,14 +22304,14 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5347_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5350_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -22323,7 +22323,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -22345,7 +22345,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyPrepareParams___clos _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5347_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5350_), 1, 0); return x_1; } } @@ -22381,7 +22381,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -22403,7 +22403,7 @@ x_6 = lean_array_get_size(x_5); x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____spec__2(x_7, x_8, x_5); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____spec__2(x_7, x_8, x_5); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -22467,7 +22467,7 @@ return x_25; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1() { _start: { lean_object* x_1; @@ -22475,78 +22475,78 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyItem", 17); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9() { _start: { lean_object* x_1; @@ -22554,67 +22554,67 @@ x_1 = lean_mk_string_from_bytes("tags\?", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16() { _start: { lean_object* x_1; @@ -22622,111 +22622,111 @@ x_1 = lean_mk_string_from_bytes("uri", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__26() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -22736,7 +22736,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -22748,7 +22748,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -22762,8 +22762,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -22773,7 +22773,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -22785,7 +22785,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -22799,8 +22799,8 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; -x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1(x_1, x_24); +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; +x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { uint8_t x_26; @@ -22811,7 +22811,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -22823,7 +22823,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -22837,7 +22837,7 @@ lean_object* x_34; lean_object* x_35; lean_object* x_36; x_34 = lean_ctor_get(x_25, 0); lean_inc(x_34); lean_dec(x_25); -x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_35 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_36 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_1, x_35); if (lean_obj_tag(x_36) == 0) { @@ -22850,7 +22850,7 @@ if (x_37 == 0) { lean_object* x_38; lean_object* x_39; lean_object* x_40; x_38 = lean_ctor_get(x_36, 0); -x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15; +x_39 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15; x_40 = lean_string_append(x_39, x_38); lean_dec(x_38); lean_ctor_set(x_36, 0, x_40); @@ -22862,7 +22862,7 @@ lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; x_41 = lean_ctor_get(x_36, 0); lean_inc(x_41); lean_dec(x_36); -x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15; +x_42 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15; x_43 = lean_string_append(x_42, x_41); lean_dec(x_41); x_44 = lean_alloc_ctor(0, 1, 0); @@ -22876,7 +22876,7 @@ lean_object* x_45; lean_object* x_46; lean_object* x_47; x_45 = lean_ctor_get(x_36, 0); lean_inc(x_45); lean_dec(x_36); -x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16; +x_46 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16; x_47 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__1(x_1, x_46); if (lean_obj_tag(x_47) == 0) { @@ -22890,7 +22890,7 @@ if (x_48 == 0) { lean_object* x_49; lean_object* x_50; lean_object* x_51; x_49 = lean_ctor_get(x_47, 0); -x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20; +x_50 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20; x_51 = lean_string_append(x_50, x_49); lean_dec(x_49); lean_ctor_set(x_47, 0, x_51); @@ -22902,7 +22902,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; x_52 = lean_ctor_get(x_47, 0); lean_inc(x_52); lean_dec(x_47); -x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20; +x_53 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20; x_54 = lean_string_append(x_53, x_52); lean_dec(x_52); x_55 = lean_alloc_ctor(0, 1, 0); @@ -22916,7 +22916,7 @@ lean_object* x_56; lean_object* x_57; lean_object* x_58; x_56 = lean_ctor_get(x_47, 0); lean_inc(x_56); lean_dec(x_47); -x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_57 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_58 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_1, x_57); if (lean_obj_tag(x_58) == 0) { @@ -22931,7 +22931,7 @@ if (x_59 == 0) { lean_object* x_60; lean_object* x_61; lean_object* x_62; x_60 = lean_ctor_get(x_58, 0); -x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22; +x_61 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22; x_62 = lean_string_append(x_61, x_60); lean_dec(x_60); lean_ctor_set(x_58, 0, x_62); @@ -22943,7 +22943,7 @@ lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; x_63 = lean_ctor_get(x_58, 0); lean_inc(x_63); lean_dec(x_58); -x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22; +x_64 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22; x_65 = lean_string_append(x_64, x_63); lean_dec(x_63); x_66 = lean_alloc_ctor(0, 1, 0); @@ -22957,7 +22957,7 @@ lean_object* x_67; lean_object* x_68; lean_object* x_69; x_67 = lean_ctor_get(x_58, 0); lean_inc(x_67); lean_dec(x_58); -x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_68 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_69 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_1, x_68); if (lean_obj_tag(x_69) == 0) { @@ -22973,7 +22973,7 @@ if (x_70 == 0) { lean_object* x_71; lean_object* x_72; lean_object* x_73; x_71 = lean_ctor_get(x_69, 0); -x_72 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24; +x_72 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24; x_73 = lean_string_append(x_72, x_71); lean_dec(x_71); lean_ctor_set(x_69, 0, x_73); @@ -22985,7 +22985,7 @@ lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; x_74 = lean_ctor_get(x_69, 0); lean_inc(x_74); lean_dec(x_69); -x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24; +x_75 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24; x_76 = lean_string_append(x_75, x_74); lean_dec(x_74); x_77 = lean_alloc_ctor(0, 1, 0); @@ -22999,8 +22999,8 @@ lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; x_78 = lean_ctor_get(x_69, 0); lean_inc(x_78); lean_dec(x_69); -x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; -x_80 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____spec__4(x_1, x_79); +x_79 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; +x_80 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____spec__4(x_1, x_79); x_81 = !lean_is_exclusive(x_80); if (x_81 == 0) { @@ -23050,21 +23050,21 @@ return x_88; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472_(x_1); lean_dec(x_1); return x_2; } @@ -23073,7 +23073,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____boxed), 1, 0); return x_1; } } @@ -23085,7 +23085,7 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -23105,7 +23105,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5155____spec__1(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5158____spec__1(x_6, x_7, x_4); x_9 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_9, 0, x_8); x_10 = lean_alloc_ctor(0, 2, 0); @@ -23119,7 +23119,7 @@ return x_12; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; @@ -23141,7 +23141,7 @@ lean_inc(x_9); lean_dec(x_1); x_10 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_10, 0, x_2); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -23149,14 +23149,14 @@ x_13 = lean_box(0); x_14 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9; -x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803____spec__1(x_15, x_4); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9; +x_16 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806____spec__1(x_15, x_4); +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10; x_18 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_17, x_5); lean_dec(x_5); x_19 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_19, 0, x_6); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16; x_21 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); @@ -23164,7 +23164,7 @@ x_22 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_22, 0, x_21); lean_ctor_set(x_22, 1, x_13); x_23 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_7); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_25 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); @@ -23172,14 +23172,14 @@ x_26 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_13); x_27 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_8); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20; x_29 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_27); x_30 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_30, 0, x_29); lean_ctor_set(x_30, 1, x_13); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; x_32 = l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__2(x_31, x_9); lean_dec(x_9); x_33 = lean_alloc_ctor(1, 2, 0); @@ -23387,7 +23387,7 @@ goto block_47; block_47: { lean_object* x_40; 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 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; +x_40 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; x_41 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_39); @@ -23410,7 +23410,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_), 1, 0); return x_1; } } @@ -23422,7 +23422,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____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; @@ -23448,7 +23448,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23498,7 +23498,7 @@ else { lean_object* x_12; uint8_t x_13; x_12 = lean_unsigned_to_nat(0u); -x_13 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__2(x_6, x_7, lean_box(0), x_6, x_7, x_12); +x_13 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__2(x_6, x_7, lean_box(0), x_6, x_7, x_12); lean_dec(x_6); return x_13; } @@ -23506,7 +23506,7 @@ return x_13; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23543,7 +23543,7 @@ return x_8; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -23580,7 +23580,7 @@ return x_8; } } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; 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; uint8_t x_19; @@ -23625,7 +23625,7 @@ return x_20; else { uint8_t x_21; -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3796_(x_4, x_12); +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3799_(x_4, x_12); if (x_21 == 0) { uint8_t x_22; @@ -23641,7 +23641,7 @@ 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__1(x_5, x_13); +x_23 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__1(x_5, x_13); if (x_23 == 0) { uint8_t x_24; @@ -23656,7 +23656,7 @@ return x_24; else { uint8_t x_25; -x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3(x_6, x_14); +x_25 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3(x_6, x_14); lean_dec(x_6); if (x_25 == 0) { @@ -23710,7 +23710,7 @@ return x_32; else { uint8_t x_33; -x_33 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4(x_10, x_18); +x_33 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4(x_10, x_18); lean_dec(x_10); return x_33; } @@ -23722,11 +23722,11 @@ return x_33; } } } -LEAN_EXPORT lean_object* l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__2(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Array_isEqvAux___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____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); @@ -23735,43 +23735,43 @@ x_8 = lean_box(x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__1(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3___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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__3(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__3(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_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4___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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4___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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____spec__4(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____spec__4(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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912_(x_1, x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915_(x_1, x_2); lean_dec(x_2); x_4 = lean_box(x_3); return x_4; @@ -23781,7 +23781,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqCallHierarchyItem___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5912____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5915____boxed), 2, 0); return x_1; } } @@ -23793,7 +23793,7 @@ x_1 = l_Lean_Lsp_instBEqCallHierarchyItem___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__1(uint64_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint64_t x_5) { +LEAN_EXPORT uint64_t l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1(uint64_t x_1, lean_object* x_2, size_t x_3, size_t x_4, uint64_t x_5) { _start: { uint8_t x_6; @@ -23814,7 +23814,7 @@ return x_5; } } } -static uint64_t _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1() { +static uint64_t _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1() { _start: { uint64_t x_1; uint64_t x_2; uint64_t x_3; @@ -23824,7 +23824,7 @@ x_3 = lean_uint64_mix_hash(x_1, x_2); return x_3; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085_(lean_object* x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint64_t x_10; uint64_t x_11; uint64_t x_12; uint64_t x_13; uint64_t x_14; uint64_t x_15; uint64_t x_16; uint64_t x_17; uint64_t x_18; @@ -23839,7 +23839,7 @@ x_9 = lean_ctor_get(x_1, 6); x_10 = 0; x_11 = lean_string_hash(x_2); x_12 = lean_uint64_mix_hash(x_10, x_11); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3812_(x_3); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSymbolKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3815_(x_3); x_14 = lean_uint64_mix_hash(x_12, x_13); x_15 = lean_string_hash(x_6); x_16 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_hashRange____x40_Lean_Data_Lsp_Basic___hyg_674_(x_7); @@ -23863,7 +23863,7 @@ if (x_51 == 0) { uint64_t x_52; lean_dec(x_49); -x_52 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1; +x_52 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1; x_18 = x_52; goto block_45; } @@ -23876,7 +23876,7 @@ if (x_54 == 0) { uint64_t x_55; lean_dec(x_49); -x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1; +x_55 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1; x_18 = x_55; goto block_45; } @@ -23886,7 +23886,7 @@ size_t x_56; size_t x_57; uint64_t x_58; uint64_t x_59; x_56 = 0; x_57 = lean_usize_of_nat(x_49); lean_dec(x_49); -x_58 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__1(x_10, x_47, x_56, x_57, x_48); +x_58 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1(x_10, x_47, x_56, x_57, x_48); x_59 = lean_uint64_mix_hash(x_58, x_53); x_18 = x_59; goto block_45; @@ -23953,7 +23953,7 @@ return x_44; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__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_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { uint64_t x_6; size_t x_7; size_t x_8; uint64_t x_9; uint64_t x_10; lean_object* x_11; @@ -23965,17 +23965,17 @@ x_8 = lean_unbox_usize(x_4); lean_dec(x_4); x_9 = lean_unbox_uint64(x_5); lean_dec(x_5); -x_10 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____spec__1(x_6, x_2, x_7, x_8, x_9); +x_10 = l_Array_foldlMUnsafe_fold___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____spec__1(x_6, x_2, x_7, x_8, x_9); lean_dec(x_2); x_11 = lean_box_uint64(x_10); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____boxed(lean_object* x_1) { _start: { uint64_t x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088_(x_1); lean_dec(x_1); x_3 = lean_box_uint64(x_2); return x_3; @@ -23985,7 +23985,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableCallHierarchyItem___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____boxed), 1, 0); return x_1; } } @@ -24047,17 +24047,17 @@ x_1 = l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472_(x_3); lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1() { _start: { lean_object* x_1; @@ -24065,7 +24065,7 @@ x_1 = lean_mk_string_from_bytes("item", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2() { _start: { lean_object* x_1; @@ -24073,83 +24073,83 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyIncomingCallsParams", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24158,7 +24158,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24170,7 +24170,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24199,21 +24199,21 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209_(x_1); lean_dec(x_1); return x_2; } @@ -24222,7 +24222,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParam _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____boxed), 1, 0); return x_1; } } @@ -24234,12 +24234,12 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6267_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6270_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(x_1); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -24259,7 +24259,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6267_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6270_), 1, 0); return x_1; } } @@ -24271,7 +24271,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -24327,7 +24327,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -24344,7 +24344,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__2(x_6, x_7, x_4); return x_8; } else @@ -24363,7 +24363,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1() { _start: { lean_object* x_1; @@ -24371,7 +24371,7 @@ x_1 = lean_mk_string_from_bytes("from", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2() { _start: { lean_object* x_1; @@ -24379,78 +24379,78 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyIncomingCall", 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10() { _start: { lean_object* x_1; @@ -24458,52 +24458,52 @@ x_1 = lean_mk_string_from_bytes("fromRanges", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24513,7 +24513,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24525,7 +24525,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24539,8 +24539,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -24550,7 +24550,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -24562,7 +24562,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -24601,7 +24601,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____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; @@ -24609,7 +24609,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__2(x_4, x_5, x_3); return x_6; } } @@ -24617,7 +24617,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322_), 1, 0); return x_1; } } @@ -24629,7 +24629,7 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -24654,14 +24654,14 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422_(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; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -24676,10 +24676,10 @@ x_9 = lean_array_get_size(x_8); x_10 = lean_usize_of_nat(x_9); lean_dec(x_9); x_11 = 0; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(x_10, x_11, x_8); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1(x_10, x_11, x_8); x_13 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -24697,7 +24697,7 @@ x_20 = l_Lean_Json_mkObj(x_19); return x_20; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____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; @@ -24705,7 +24705,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1(x_4, x_5, x_3); return x_6; } } @@ -24713,7 +24713,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCall___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422_), 1, 0); return x_1; } } @@ -24745,7 +24745,7 @@ x_1 = l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1() { _start: { lean_object* x_1; @@ -24753,63 +24753,63 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyOutgoingCallsParams", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -24818,7 +24818,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -24830,7 +24830,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -24859,11 +24859,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489_(x_1); lean_dec(x_1); return x_2; } @@ -24872,7 +24872,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParam _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____boxed), 1, 0); return x_1; } } @@ -24884,12 +24884,12 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6547_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6550_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(x_1); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -24909,7 +24909,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6547_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6550_), 1, 0); return x_1; } } @@ -24921,7 +24921,7 @@ x_1 = l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1() { _start: { lean_object* x_1; @@ -24929,7 +24929,7 @@ x_1 = lean_mk_string_from_bytes("to", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2() { _start: { lean_object* x_1; @@ -24937,103 +24937,103 @@ x_1 = lean_mk_string_from_bytes("CallHierarchyOutgoingCall", 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -25043,7 +25043,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -25055,7 +25055,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -25069,8 +25069,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -25080,7 +25080,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -25092,7 +25092,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -25135,7 +25135,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602_), 1, 0); return x_1; } } @@ -25147,14 +25147,14 @@ x_1 = l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6699_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6702_(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; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5803_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5806_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -25169,10 +25169,10 @@ x_9 = lean_array_get_size(x_8); x_10 = lean_usize_of_nat(x_9); lean_dec(x_9); x_11 = 0; -x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6419____spec__1(x_10, x_11, x_8); +x_12 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6422____spec__1(x_10, x_11, x_8); x_13 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_13, 0, x_12); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10; x_15 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_13); @@ -25194,7 +25194,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCall___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6699_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6702_), 1, 0); return x_1; } } @@ -25415,7 +25415,7 @@ x_6 = l_Lean_Lsp_SemanticTokenType_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1() { _start: { lean_object* x_1; @@ -25423,17 +25423,17 @@ x_1 = lean_mk_string_from_bytes("keyword", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3() { _start: { lean_object* x_1; @@ -25441,17 +25441,17 @@ x_1 = lean_mk_string_from_bytes("variable", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5() { _start: { lean_object* x_1; @@ -25459,17 +25459,17 @@ x_1 = lean_mk_string_from_bytes("property", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7() { _start: { lean_object* x_1; @@ -25477,17 +25477,17 @@ x_1 = lean_mk_string_from_bytes("function", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9() { _start: { lean_object* x_1; @@ -25495,17 +25495,17 @@ x_1 = lean_mk_string_from_bytes("namespace", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11() { _start: { lean_object* x_1; @@ -25513,17 +25513,17 @@ x_1 = lean_mk_string_from_bytes("type", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13() { _start: { lean_object* x_1; @@ -25531,17 +25531,17 @@ x_1 = lean_mk_string_from_bytes("class", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15() { _start: { lean_object* x_1; @@ -25549,17 +25549,17 @@ x_1 = lean_mk_string_from_bytes("enum", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17() { _start: { lean_object* x_1; @@ -25567,17 +25567,17 @@ x_1 = lean_mk_string_from_bytes("interface", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19() { _start: { lean_object* x_1; @@ -25585,17 +25585,17 @@ x_1 = lean_mk_string_from_bytes("struct", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21() { _start: { lean_object* x_1; @@ -25603,17 +25603,17 @@ x_1 = lean_mk_string_from_bytes("typeParameter", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23() { _start: { lean_object* x_1; @@ -25621,17 +25621,17 @@ x_1 = lean_mk_string_from_bytes("parameter", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25() { _start: { lean_object* x_1; @@ -25639,17 +25639,17 @@ x_1 = lean_mk_string_from_bytes("enumMember", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27() { _start: { lean_object* x_1; @@ -25657,17 +25657,17 @@ x_1 = lean_mk_string_from_bytes("event", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29() { _start: { lean_object* x_1; @@ -25675,17 +25675,17 @@ x_1 = lean_mk_string_from_bytes("method", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31() { _start: { lean_object* x_1; @@ -25693,17 +25693,17 @@ x_1 = lean_mk_string_from_bytes("macro", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33() { _start: { lean_object* x_1; @@ -25711,17 +25711,17 @@ x_1 = lean_mk_string_from_bytes("modifier", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35() { _start: { lean_object* x_1; @@ -25729,17 +25729,17 @@ x_1 = lean_mk_string_from_bytes("comment", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37() { _start: { lean_object* x_1; @@ -25747,17 +25747,17 @@ x_1 = lean_mk_string_from_bytes("string", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39() { _start: { lean_object* x_1; @@ -25765,17 +25765,17 @@ x_1 = lean_mk_string_from_bytes("number", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41() { _start: { lean_object* x_1; @@ -25783,17 +25783,17 @@ x_1 = lean_mk_string_from_bytes("regexp", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43() { _start: { lean_object* x_1; @@ -25801,17 +25801,17 @@ x_1 = lean_mk_string_from_bytes("operator", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45() { _start: { lean_object* x_1; @@ -25819,17 +25819,17 @@ x_1 = lean_mk_string_from_bytes("decorator", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47() { _start: { lean_object* x_1; @@ -25837,174 +25837,174 @@ x_1 = lean_mk_string_from_bytes("leanSorryLike", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762_(uint8_t x_1) { _start: { switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2; return x_2; } case 1: { lean_object* x_3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4; return x_3; } case 2: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6; return x_4; } case 3: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8; return x_5; } case 4: { lean_object* x_6; -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10; return x_6; } case 5: { lean_object* x_7; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12; return x_7; } case 6: { lean_object* x_8; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14; return x_8; } case 7: { lean_object* x_9; -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16; return x_9; } case 8: { lean_object* x_10; -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18; return x_10; } case 9: { lean_object* x_11; -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20; return x_11; } case 10: { lean_object* x_12; -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22; return x_12; } case 11: { lean_object* x_13; -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24; return x_13; } case 12: { lean_object* x_14; -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26; return x_14; } case 13: { lean_object* x_15; -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28; return x_15; } case 14: { lean_object* x_16; -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30; return x_16; } case 15: { lean_object* x_17; -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32; return x_17; } case 16: { lean_object* x_18; -x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34; +x_18 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34; return x_18; } case 17: { lean_object* x_19; -x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36; +x_19 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36; return x_19; } case 18: { lean_object* x_20; -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38; return x_20; } case 19: { lean_object* x_21; -x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40; +x_21 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40; return x_21; } case 20: { lean_object* x_22; -x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42; +x_22 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42; return x_22; } case 21: { lean_object* x_23; -x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44; +x_23 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44; return x_23; } case 22: { lean_object* x_24; -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46; return x_24; } default: { lean_object* x_25; -x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48; +x_25 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48; return x_25; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762_(x_2); return x_3; } } @@ -26012,7 +26012,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokenType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____boxed), 1, 0); return x_1; } } @@ -26024,7 +26024,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokenType___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -26032,33 +26032,33 @@ x_1 = lean_mk_string_from_bytes("no inductive constructor matched", 32); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2; return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___boxed), 1, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26069,21 +26069,21 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_3 = l_Except_orElseLazy___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); if (lean_obj_tag(x_6) == 0) @@ -26093,7 +26093,7 @@ x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_9 = l_Except_orElseLazy___rarg(x_6, x_8); return x_9; } @@ -26105,7 +26105,7 @@ lean_inc(x_10); lean_dec(x_6); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_13 = l_Except_orElseLazy___rarg(x_11, x_12); return x_13; } @@ -26114,12 +26114,12 @@ else { lean_object* x_14; lean_dec(x_6); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3; return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26130,14 +26130,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26166,13 +26166,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26183,14 +26183,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26219,13 +26219,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26236,14 +26236,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26272,13 +26272,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26289,14 +26289,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26325,13 +26325,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26342,14 +26342,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26378,13 +26378,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26395,14 +26395,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26431,13 +26431,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26448,14 +26448,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__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; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26484,13 +26484,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26501,14 +26501,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26537,13 +26537,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26554,14 +26554,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26590,13 +26590,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26607,14 +26607,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26643,13 +26643,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26660,14 +26660,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26696,13 +26696,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26713,14 +26713,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26749,13 +26749,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26766,14 +26766,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26802,13 +26802,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26819,14 +26819,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26855,13 +26855,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26872,14 +26872,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26908,13 +26908,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26925,14 +26925,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -26961,13 +26961,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -26978,14 +26978,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27014,13 +27014,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27031,14 +27031,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27067,13 +27067,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27084,14 +27084,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27120,13 +27120,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27137,14 +27137,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27173,13 +27173,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27190,14 +27190,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27226,13 +27226,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27243,14 +27243,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -27279,13 +27279,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -27295,7 +27295,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -27306,15 +27306,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921_(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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47; x_3 = lean_unsigned_to_nat(0u); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1; x_5 = l_Lean_Json_parseTagged(x_1, x_2, x_3, x_4); -x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___boxed), 3, 2); +x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___boxed), 3, 2); lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_4); if (lean_obj_tag(x_5) == 0) @@ -27343,226 +27343,226 @@ else { lean_object* x_12; lean_object* x_13; lean_dec(x_5); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2; x_13 = l_Except_orElseLazy___rarg(x_12, x_6); return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1(x_1); lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -27571,7 +27571,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921_), 1, 0); return x_1; } } @@ -27583,7 +27583,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -27595,7 +27595,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751____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; @@ -27603,7 +27603,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748_(x_3, x_4); +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -27612,7 +27612,7 @@ static lean_object* _init_l_Lean_Lsp_instBEqSemanticTokenType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7748____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_beqSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7751____boxed), 2, 0); return x_1; } } @@ -27624,7 +27624,7 @@ x_1 = l_Lean_Lsp_instBEqSemanticTokenType___closed__1; return x_1; } } -LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764_(uint8_t x_1) { +LEAN_EXPORT uint64_t l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767_(uint8_t x_1) { _start: { switch (x_1) { @@ -27775,13 +27775,13 @@ return x_25; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767____boxed(lean_object* x_1) { _start: { uint8_t x_2; uint64_t x_3; lean_object* x_4; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767_(x_2); x_4 = lean_box_uint64(x_3); return x_4; } @@ -27790,7 +27790,7 @@ static lean_object* _init_l_Lean_Lsp_instHashableSemanticTokenType___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7764____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7767____boxed), 1, 0); return x_1; } } @@ -27816,7 +27816,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__2() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__1; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27826,7 +27826,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27836,7 +27836,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27846,7 +27846,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27856,7 +27856,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27866,7 +27866,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27876,7 +27876,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__7; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27886,7 +27886,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__8; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27896,7 +27896,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__10() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__9; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27906,7 +27906,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__11() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__10; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27916,7 +27916,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__12() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__11; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27926,7 +27926,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__13() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__12; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27936,7 +27936,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__14() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__13; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27946,7 +27946,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__15() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__14; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27956,7 +27956,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__16() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__15; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27966,7 +27966,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__17() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__16; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27976,7 +27976,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__18() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__17; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27986,7 +27986,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__19() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__18; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -27996,7 +27996,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__20() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__19; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28006,7 +28006,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__21() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__20; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28016,7 +28016,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__22() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__21; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28026,7 +28026,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__23() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__22; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28036,7 +28036,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__24() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__23; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28046,7 +28046,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenType_names___closed__25() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenType_names___closed__24; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -28182,7 +28182,7 @@ x_6 = l_Lean_Lsp_SemanticTokenModifier_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1() { _start: { lean_object* x_1; @@ -28190,17 +28190,17 @@ x_1 = lean_mk_string_from_bytes("declaration", 11); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3() { _start: { lean_object* x_1; @@ -28208,17 +28208,17 @@ x_1 = lean_mk_string_from_bytes("definition", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5() { _start: { lean_object* x_1; @@ -28226,17 +28226,17 @@ x_1 = lean_mk_string_from_bytes("readonly", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7() { _start: { lean_object* x_1; @@ -28244,17 +28244,17 @@ x_1 = lean_mk_string_from_bytes("static", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9() { _start: { lean_object* x_1; @@ -28262,17 +28262,17 @@ x_1 = lean_mk_string_from_bytes("deprecated", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11() { _start: { lean_object* x_1; @@ -28280,17 +28280,17 @@ x_1 = lean_mk_string_from_bytes("abstract", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13() { _start: { lean_object* x_1; @@ -28298,17 +28298,17 @@ x_1 = lean_mk_string_from_bytes("async", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15() { _start: { lean_object* x_1; @@ -28316,27 +28316,27 @@ x_1 = lean_mk_string_from_bytes("modification", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18() { _start: { lean_object* x_1; @@ -28344,90 +28344,90 @@ x_1 = lean_mk_string_from_bytes("defaultLibrary", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18; 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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978_(uint8_t x_1) { _start: { switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2; return x_2; } case 1: { lean_object* x_3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4; return x_3; } case 2: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6; return x_4; } case 3: { lean_object* x_5; -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8; return x_5; } case 4: { lean_object* x_6; -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10; return x_6; } case 5: { lean_object* x_7; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12; return x_7; } case 6: { lean_object* x_8; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14; return x_8; } case 7: { lean_object* x_9; -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16; return x_9; } case 8: { lean_object* x_10; -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17; return x_10; } default: { lean_object* x_11; -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19; return x_11; } } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978_(x_2); return x_3; } } @@ -28435,7 +28435,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____boxed), 1, 0); return x_1; } } @@ -28447,7 +28447,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28458,21 +28458,21 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_3 = l_Except_orElseLazy___rarg(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____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 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); if (lean_obj_tag(x_6) == 0) @@ -28482,7 +28482,7 @@ x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { lean_object* x_8; lean_object* x_9; -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_9 = l_Except_orElseLazy___rarg(x_6, x_8); return x_9; } @@ -28494,7 +28494,7 @@ lean_inc(x_10); lean_dec(x_6); x_11 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1; x_13 = l_Except_orElseLazy___rarg(x_11, x_12); return x_13; } @@ -28503,12 +28503,12 @@ else { lean_object* x_14; lean_dec(x_6); -x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2; +x_14 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2; return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28519,14 +28519,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28555,13 +28555,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28572,14 +28572,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28608,13 +28608,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28625,14 +28625,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28661,13 +28661,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28678,14 +28678,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28714,13 +28714,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28731,14 +28731,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28767,13 +28767,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28784,14 +28784,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28820,13 +28820,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28837,14 +28837,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28873,13 +28873,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28890,14 +28890,14 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__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; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1; x_5 = lean_unsigned_to_nat(0u); x_6 = l_Lean_Json_parseTagged(x_1, x_4, x_5, x_2); -x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___boxed), 3, 2); +x_7 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___boxed), 3, 2); lean_closure_set(x_7, 0, x_1); lean_closure_set(x_7, 1, x_2); if (lean_obj_tag(x_6) == 0) @@ -28926,13 +28926,13 @@ else { lean_object* x_13; lean_object* x_14; lean_dec(x_6); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1; x_14 = l_Except_orElseLazy___rarg(x_13, x_7); return x_14; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; @@ -28943,15 +28943,15 @@ lean_ctor_set(x_3, 0, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053_(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___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18; x_3 = lean_unsigned_to_nat(0u); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1; x_5 = l_Lean_Json_parseTagged(x_1, x_2, x_3, x_4); -x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___boxed), 3, 2); +x_6 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___boxed), 3, 2); lean_closure_set(x_6, 0, x_1); lean_closure_set(x_6, 1, x_4); if (lean_obj_tag(x_5) == 0) @@ -28980,91 +28980,91 @@ else { lean_object* x_12; lean_object* x_13; lean_dec(x_5); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1; x_13 = l_Except_orElseLazy___rarg(x_12, x_6); return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8(x_1, x_2, x_3); lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9(x_1, x_2, x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9(x_1, x_2, x_3); lean_dec(x_3); return x_4; } @@ -29073,7 +29073,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053_), 1, 0); return x_1; } } @@ -29099,7 +29099,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__2() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__1; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29109,7 +29109,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__3() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29119,7 +29119,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__4() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29129,7 +29129,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__5() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29139,7 +29139,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__6() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29149,7 +29149,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__7() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__6; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29159,7 +29159,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__8() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__7; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29169,7 +29169,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__9() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__8; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29179,7 +29179,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__10() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__9; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29189,7 +29189,7 @@ static lean_object* _init_l_Lean_Lsp_SemanticTokenModifier_names___closed__11() { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Lsp_SemanticTokenModifier_names___closed__10; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18; x_3 = lean_array_push(x_1, x_2); return x_3; } @@ -29220,7 +29220,7 @@ x_3 = l_Lean_Lsp_SemanticTokenModifier_toNat(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -29254,7 +29254,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1() { _start: { lean_object* x_1; @@ -29262,7 +29262,7 @@ x_1 = lean_mk_string_from_bytes("tokenTypes", 10); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2() { _start: { lean_object* x_1; @@ -29270,78 +29270,78 @@ x_1 = lean_mk_string_from_bytes("SemanticTokensLegend", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10() { _start: { lean_object* x_1; @@ -29349,52 +29349,52 @@ x_1 = lean_mk_string_from_bytes("tokenModifiers", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -29403,7 +29403,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -29415,7 +29415,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -29429,8 +29429,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -29440,7 +29440,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -29452,7 +29452,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -29491,21 +29491,21 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554_(x_1); lean_dec(x_1); return x_2; } @@ -29514,7 +29514,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____boxed), 1, 0); return x_1; } } @@ -29526,7 +29526,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8651_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8654_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; size_t x_4; size_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; size_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -29539,7 +29539,7 @@ x_5 = 0; x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182____spec__2(x_4, x_5, x_2); x_7 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_7, 0, x_6); -x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1; +x_8 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1; x_9 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_9, 0, x_8); lean_ctor_set(x_9, 1, x_7); @@ -29556,7 +29556,7 @@ lean_dec(x_13); x_15 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_182____spec__2(x_14, x_5, x_12); x_16 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_16, 0, x_15); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10; x_18 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); @@ -29578,7 +29578,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8651_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8654_), 1, 0); return x_1; } } @@ -29590,17 +29590,17 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; x_3 = l_Lean_Json_getObjValD(x_1, x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551_(x_3); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554_(x_3); lean_dec(x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1() { _start: { lean_object* x_1; @@ -29608,7 +29608,7 @@ x_1 = lean_mk_string_from_bytes("legend", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2() { _start: { lean_object* x_1; @@ -29616,98 +29616,98 @@ x_1 = lean_mk_string_from_bytes("SemanticTokensOptions", 21); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12() { _start: { lean_object* x_1; @@ -29715,52 +29715,52 @@ x_1 = lean_mk_string_from_bytes("full", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1(x_1, x_2); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1; +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { uint8_t x_4; @@ -29769,7 +29769,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -29781,7 +29781,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -29795,7 +29795,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -29806,7 +29806,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -29818,7 +29818,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -29832,7 +29832,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12; x_25 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -29844,7 +29844,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -29856,7 +29856,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -29906,21 +29906,21 @@ return x_43; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____spec__1(x_1, x_2); +x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____spec__1(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727_(x_1); lean_dec(x_1); return x_2; } @@ -29929,7 +29929,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____boxed), 1, 0); return x_1; } } @@ -29941,14 +29941,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8863_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8866_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8651_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8654_(x_2); +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -29959,7 +29959,7 @@ lean_ctor_set(x_7, 1, x_6); x_8 = lean_ctor_get_uint8(x_1, sizeof(void*)*1); x_9 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_9, 0, x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -29970,7 +29970,7 @@ x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*1 + 1); lean_dec(x_1); x_14 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_14, 0, x_13); -x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12; +x_15 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12; x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); lean_ctor_set(x_16, 1, x_14); @@ -29995,7 +29995,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8863_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8866_), 1, 0); return x_1; } } @@ -30007,7 +30007,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1() { _start: { lean_object* x_1; @@ -30015,62 +30015,62 @@ x_1 = lean_mk_string_from_bytes("SemanticTokensParams", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -30080,7 +30080,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -30092,7 +30092,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -30121,11 +30121,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937_(x_1); lean_dec(x_1); return x_2; } @@ -30134,7 +30134,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____boxed), 1, 0); return x_1; } } @@ -30146,12 +30146,12 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8995_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8998_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -30171,7 +30171,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8995_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8998_), 1, 0); return x_1; } } @@ -30183,7 +30183,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1() { _start: { lean_object* x_1; @@ -30191,82 +30191,82 @@ x_1 = lean_mk_string_from_bytes("SemanticTokensRangeParams", 25); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -30276,7 +30276,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -30288,7 +30288,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -30302,7 +30302,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonLocation____x40_Lean_Data_Lsp_Basic___hyg_1084____spec__2(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -30313,7 +30313,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -30325,7 +30325,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -30364,11 +30364,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049_(x_1); lean_dec(x_1); return x_2; } @@ -30377,7 +30377,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___clo _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____boxed), 1, 0); return x_1; } } @@ -30389,14 +30389,14 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9146_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -30408,7 +30408,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonRange____x40_Lean_Data_Lsp_Basic___hyg_714_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -30430,7 +30430,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokensRangeParams___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9146_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9149_), 1, 0); return x_1; } } @@ -30450,7 +30450,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__2(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -30506,7 +30506,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; @@ -30523,7 +30523,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2(x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__2(x_6, x_7, x_4); return x_8; } else @@ -30542,7 +30542,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1() { _start: { lean_object* x_1; @@ -30550,7 +30550,7 @@ x_1 = lean_mk_string_from_bytes("resultId", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2() { _start: { lean_object* x_1; @@ -30558,38 +30558,38 @@ x_1 = lean_mk_string_from_bytes("SemanticTokens", 14); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6() { _start: { lean_object* x_1; @@ -30597,91 +30597,91 @@ x_1 = lean_mk_string_from_bytes("resultId\?", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____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_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__7; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1956____spec__2(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -30692,7 +30692,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -30704,7 +30704,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -30718,8 +30718,8 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; -x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__1(x_1, x_13); +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; +x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { uint8_t x_15; @@ -30729,7 +30729,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -30741,7 +30741,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -30780,7 +30780,7 @@ return x_28; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____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; @@ -30788,7 +30788,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____spec__2(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____spec__2(x_4, x_5, x_3); return x_6; } } @@ -30796,7 +30796,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217_), 1, 0); return x_1; } } @@ -30808,7 +30808,7 @@ x_1 = l_Lean_Lsp_instFromJsonSemanticTokens___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____spec__1(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -30835,13 +30835,13 @@ goto _start; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317_(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; size_t x_7; size_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1; x_4 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextEdit____x40_Lean_Data_Lsp_Basic___hyg_1896____spec__2(x_3, x_2); lean_dec(x_2); x_5 = lean_ctor_get(x_1, 1); @@ -30851,10 +30851,10 @@ x_6 = lean_array_get_size(x_5); x_7 = lean_usize_of_nat(x_6); lean_dec(x_6); x_8 = 0; -x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1(x_7, x_8, x_5); +x_9 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____spec__1(x_7, x_8, x_5); x_10 = lean_alloc_ctor(4, 1, 0); lean_ctor_set(x_10, 0, x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -30873,7 +30873,7 @@ x_18 = l_Lean_Json_mkObj(x_17); return x_18; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____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; @@ -30881,7 +30881,7 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314____spec__1(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317____spec__1(x_4, x_5, x_3); return x_6; } } @@ -30889,7 +30889,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonSemanticTokens___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9314_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9317_), 1, 0); return x_1; } } @@ -30901,7 +30901,7 @@ x_1 = l_Lean_Lsp_instToJsonSemanticTokens___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1() { _start: { lean_object* x_1; @@ -30909,62 +30909,62 @@ x_1 = lean_mk_string_from_bytes("FoldingRangeParams", 18); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -30974,7 +30974,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -30986,7 +30986,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31015,11 +31015,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367_(x_1); lean_dec(x_1); return x_2; } @@ -31028,7 +31028,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1( _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____boxed), 1, 0); return x_1; } } @@ -31040,12 +31040,12 @@ x_1 = l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9425_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9428_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -31065,7 +31065,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRangeParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9425_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9428_), 1, 0); return x_1; } } @@ -31183,7 +31183,7 @@ switch (x_1) { case 0: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36; return x_2; } case 1: @@ -31219,7 +31219,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -31239,7 +31239,7 @@ switch (x_6) { case 0: { lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36; +x_7 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36; x_8 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_8, 0, x_1); lean_ctor_set(x_8, 1, x_7); @@ -31276,7 +31276,7 @@ return x_15; } } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1() { _start: { lean_object* x_1; @@ -31284,7 +31284,7 @@ x_1 = lean_mk_string_from_bytes("startLine", 9); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2() { _start: { lean_object* x_1; @@ -31292,7 +31292,7 @@ x_1 = lean_mk_string_from_bytes("endLine", 7); return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -31301,7 +31301,7 @@ lean_inc(x_2); x_3 = l_Lean_JsonNumber_fromNat(x_2); x_4 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_4, 0, x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -31314,7 +31314,7 @@ lean_inc(x_9); x_10 = l_Lean_JsonNumber_fromNat(x_9); x_11 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_11, 0, x_10); -x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2; +x_12 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2; x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); lean_ctor_set(x_13, 1, x_11); @@ -31324,8 +31324,8 @@ lean_ctor_set(x_14, 1, x_7); x_15 = lean_ctor_get(x_1, 2); lean_inc(x_15); lean_dec(x_1); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22; -x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1(x_16, x_15); +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22; +x_17 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1(x_16, x_15); lean_dec(x_15); x_18 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_18, 0, x_17); @@ -31341,11 +31341,11 @@ x_22 = l_Lean_Json_mkObj(x_21); return x_22; } } -LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____spec__1(x_1, x_2); +x_3 = l_Lean_Json_opt___at___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____spec__1(x_1, x_2); lean_dec(x_2); return x_3; } @@ -31354,7 +31354,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524_), 1, 0); return x_1; } } @@ -31374,7 +31374,7 @@ x_1 = 0; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1() { _start: { lean_object* x_1; @@ -31382,7 +31382,7 @@ x_1 = lean_mk_string_from_bytes("prepareProvider", 15); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2() { _start: { lean_object* x_1; @@ -31390,82 +31390,82 @@ x_1 = lean_mk_string_from_bytes("RenameOptions", 13); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonChangeAnnotation____x40_Lean_Data_Lsp_Basic___hyg_2710____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -31475,7 +31475,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -31487,7 +31487,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31516,11 +31516,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589_(x_1); lean_dec(x_1); return x_2; } @@ -31529,7 +31529,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameOptions___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____boxed), 1, 0); return x_1; } } @@ -31541,13 +31541,13 @@ x_1 = l_Lean_Lsp_instFromJsonRenameOptions___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647_(uint8_t x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650_(uint8_t x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; x_2 = lean_alloc_ctor(1, 0, 1); lean_ctor_set_uint8(x_2, 0, x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1; x_4 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_4, 0, x_3); lean_ctor_set(x_4, 1, x_2); @@ -31563,13 +31563,13 @@ x_9 = l_Lean_Json_mkObj(x_8); return x_9; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650____boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; x_2 = lean_unbox(x_1); lean_dec(x_1); -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647_(x_2); +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650_(x_2); return x_3; } } @@ -31577,7 +31577,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameOptions___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9647____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9650____boxed), 1, 0); return x_1; } } @@ -31589,7 +31589,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameOptions___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1() { _start: { lean_object* x_1; @@ -31597,78 +31597,78 @@ x_1 = lean_mk_string_from_bytes("RenameParams", 12); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9() { _start: { lean_object* x_1; @@ -31676,51 +31676,51 @@ x_1 = lean_mk_string_from_bytes("newName", 7); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -31730,7 +31730,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -31742,7 +31742,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -31756,7 +31756,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -31767,7 +31767,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -31779,7 +31779,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -31793,7 +31793,7 @@ lean_object* x_23; lean_object* x_24; lean_object* x_25; x_23 = lean_ctor_get(x_14, 0); lean_inc(x_23); lean_dec(x_14); -x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9; +x_24 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9; x_25 = l_Lean_Json_getObjValAs_x3f___at_Lean_JsonRpc_instFromJsonMessage___spec__3(x_1, x_24); if (lean_obj_tag(x_25) == 0) { @@ -31805,7 +31805,7 @@ if (x_26 == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; x_27 = lean_ctor_get(x_25, 0); -x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13; +x_28 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); lean_ctor_set(x_25, 0, x_29); @@ -31817,7 +31817,7 @@ lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; x_30 = lean_ctor_get(x_25, 0); lean_inc(x_30); lean_dec(x_25); -x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13; +x_31 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13; x_32 = lean_string_append(x_31, x_30); lean_dec(x_30); x_33 = lean_alloc_ctor(0, 1, 0); @@ -31863,11 +31863,11 @@ return x_41; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701_(x_1); lean_dec(x_1); return x_2; } @@ -31876,7 +31876,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonRenameParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____boxed), 1, 0); return x_1; } } @@ -31888,7 +31888,7 @@ x_1 = l_Lean_Lsp_instFromJsonRenameParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9837_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9840_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; @@ -31897,7 +31897,7 @@ lean_inc(x_2); x_3 = lean_ctor_get(x_2, 0); lean_inc(x_3); x_4 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_3); -x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_5 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_6 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_6, 0, x_5); lean_ctor_set(x_6, 1, x_4); @@ -31909,7 +31909,7 @@ x_9 = lean_ctor_get(x_2, 1); lean_inc(x_9); lean_dec(x_2); x_10 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_9); -x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_11 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_12 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_12, 0, x_11); lean_ctor_set(x_12, 1, x_10); @@ -31921,7 +31921,7 @@ lean_inc(x_14); lean_dec(x_1); x_15 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_15, 0, x_14); -x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9; +x_16 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9; x_17 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); @@ -31946,7 +31946,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonRenameParams___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9837_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9840_), 1, 0); return x_1; } } @@ -31958,7 +31958,7 @@ x_1 = l_Lean_Lsp_instToJsonRenameParams___closed__1; return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1() { _start: { lean_object* x_1; @@ -31966,82 +31966,82 @@ x_1 = lean_mk_string_from_bytes("PrepareRenameParams", 19); return x_1; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__2; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__3; -x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1; +x_3 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2; x_2 = 1; x_3 = l_Lean_Name_toString(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12; x_3 = lean_string_append(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8() { +static lean_object* _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7; +x_1 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7; x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_43____closed__13; x_3 = lean_string_append(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_3 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonTextDocumentPositionParams____x40_Lean_Data_Lsp_Basic___hyg_5082____spec__1(x_1, x_2); if (lean_obj_tag(x_3) == 0) { @@ -32051,7 +32051,7 @@ if (x_4 == 0) { lean_object* x_5; lean_object* x_6; lean_object* x_7; x_5 = lean_ctor_get(x_3, 0); -x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6; +x_6 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6; x_7 = lean_string_append(x_6, x_5); lean_dec(x_5); lean_ctor_set(x_3, 0, x_7); @@ -32063,7 +32063,7 @@ lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; x_8 = lean_ctor_get(x_3, 0); lean_inc(x_8); lean_dec(x_3); -x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6; +x_9 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6; x_10 = lean_string_append(x_9, x_8); lean_dec(x_8); x_11 = lean_alloc_ctor(0, 1, 0); @@ -32077,7 +32077,7 @@ lean_object* x_12; lean_object* x_13; lean_object* x_14; x_12 = lean_ctor_get(x_3, 0); lean_inc(x_12); lean_dec(x_3); -x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_13 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_14 = l_Lean_Json_getObjValAs_x3f___at___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_fromJsonRange____x40_Lean_Data_Lsp_Basic___hyg_760____spec__1(x_1, x_13); if (lean_obj_tag(x_14) == 0) { @@ -32088,7 +32088,7 @@ if (x_15 == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; x_16 = lean_ctor_get(x_14, 0); -x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8; +x_17 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8; x_18 = lean_string_append(x_17, x_16); lean_dec(x_16); lean_ctor_set(x_14, 0, x_18); @@ -32100,7 +32100,7 @@ lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; x_19 = lean_ctor_get(x_14, 0); lean_inc(x_19); lean_dec(x_14); -x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8; +x_20 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8; x_21 = lean_string_append(x_20, x_19); lean_dec(x_19); x_22 = lean_alloc_ctor(0, 1, 0); @@ -32139,11 +32139,11 @@ return x_28; } } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908_(x_1); +x_2 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911_(x_1); lean_dec(x_1); return x_2; } @@ -32152,7 +32152,7 @@ static lean_object* _init_l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1 _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____boxed), 1, 0); return x_1; } } @@ -32164,14 +32164,14 @@ x_1 = l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1; return x_1; } } -LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10008_(lean_object* x_1) { +LEAN_EXPORT lean_object* l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10011_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_2 = lean_ctor_get(x_1, 0); lean_inc(x_2); x_3 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonTextDocumentIdentifier____x40_Lean_Data_Lsp_Basic___hyg_2209_(x_2); -x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1; +x_4 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -32183,7 +32183,7 @@ x_8 = lean_ctor_get(x_1, 1); lean_inc(x_8); lean_dec(x_1); x_9 = l___private_Lean_Data_Lsp_Basic_0__Lean_Lsp_toJsonPosition____x40_Lean_Data_Lsp_Basic___hyg_393_(x_8); -x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10; +x_10 = l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10; x_11 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -32205,7 +32205,7 @@ static lean_object* _init_l_Lean_Lsp_instToJsonPrepareRenameParams___closed__1() _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10008_), 1, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_10011_), 1, 0); return x_1; } } @@ -32302,352 +32302,352 @@ lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionOptions); l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1 = _init_l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1(); lean_mark_persistent(l_Lean_Lsp_CompletionItemKind_noConfusion___rarg___closed__1); l_Lean_Lsp_instInhabitedCompletionItemKind = _init_l_Lean_Lsp_instInhabitedCompletionItemKind(); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__45); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__47); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__48); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__49); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__50); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__51); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__52); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__53); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__54); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__55); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__56); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__57); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__58); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__59); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__60); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__61); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__62); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__63); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__64); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__65); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__66); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__67); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__68); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__69); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__70); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__71); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__72); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__73); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__74); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__75); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__76); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__77); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__78); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__79); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__80); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__81); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__82); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__83); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__84); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__85); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__86); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__87); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__88); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__89); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__90); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__91); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__92); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__93); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__94); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__95); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__96); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__97); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__98); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__99); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__100); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__101); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__102); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__103); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__104); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__105); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__106); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__107); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__108); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__109); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__110); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__111); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__112); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__113); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__114); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__115); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__116); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__117); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__118); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__119); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__120); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__121); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__122); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__123); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__124); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__125); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__126); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__127); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__128); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__129); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__130); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__131); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__132); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__133); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__134); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__135); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__136); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__137); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__138); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__139); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__140); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__141); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__142); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__143); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__144); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__145); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__146); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__147); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__148); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__149); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__150); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__151); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_339____closed__152); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__46); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__47); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__48); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__49); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__50); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__51); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__52); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__53); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__54); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__55); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__56); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__57); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__58); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__59); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__60); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__61); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__62); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__63); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__64); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__65); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__66); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__67); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__68); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__69); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__70); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__71); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__72); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__73); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__74); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__75); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__76); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__77); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__78); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__79); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__80); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__81); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__82); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__83); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__84); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__85); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__86); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__87); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__88); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__89); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__90); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__91); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__92); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__93); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__94); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__95); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__96); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__97); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__98); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__99); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__100); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__101); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__102); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__103); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__104); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__105); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__106); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__107); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__108); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__109); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__110); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__111); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__112); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__113); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__114); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__115); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__116); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__117); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__118); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__119); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__120); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__121); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__122); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__123); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__124); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__125); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__126); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__127); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__128); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__129); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__130); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__131); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__132); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__133); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__134); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__135); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__136); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__137); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__138); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__139); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__140); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__141); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__142); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__143); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__144); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__145); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__146); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__147); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__148); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__149); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__150); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__151); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_reprCompletionItemKind____x40_Lean_Data_Lsp_LanguageFeatures___hyg_342____closed__152); l_Lean_Lsp_instReprCompletionItemKind___closed__1 = _init_l_Lean_Lsp_instReprCompletionItemKind___closed__1(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind___closed__1); l_Lean_Lsp_instReprCompletionItemKind = _init_l_Lean_Lsp_instReprCompletionItemKind(); lean_mark_persistent(l_Lean_Lsp_instReprCompletionItemKind); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1241____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonInsertReplaceEdit____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1244____closed__19); l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1 = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonInsertReplaceEdit___closed__1); l_Lean_Lsp_instFromJsonInsertReplaceEdit = _init_l_Lean_Lsp_instFromJsonInsertReplaceEdit(); @@ -32668,96 +32668,96 @@ l_Lean_Lsp_CompletionItem_sortText_x3f___default = _init_l_Lean_Lsp_CompletionIt lean_mark_persistent(l_Lean_Lsp_CompletionItem_sortText_x3f___default); l_Lean_Lsp_CompletionItem_data_x3f___default = _init_l_Lean_Lsp_CompletionItem_data_x3f___default(); lean_mark_persistent(l_Lean_Lsp_CompletionItem_data_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1511____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1514____closed__45); l_Lean_Lsp_instFromJsonCompletionItem___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionItem___closed__1); l_Lean_Lsp_instFromJsonCompletionItem = _init_l_Lean_Lsp_instFromJsonCompletionItem(); @@ -32772,34 +32772,34 @@ l_Lean_Lsp_instInhabitedCompletionItem___closed__2 = _init_l_Lean_Lsp_instInhabi lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem___closed__2); l_Lean_Lsp_instInhabitedCompletionItem = _init_l_Lean_Lsp_instInhabitedCompletionItem(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCompletionItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1915____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionList____x40_Lean_Data_Lsp_LanguageFeatures___hyg_1918____closed__14); l_Lean_Lsp_instFromJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionList___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionList___closed__1); l_Lean_Lsp_instFromJsonCompletionList = _init_l_Lean_Lsp_instFromJsonCompletionList(); @@ -32808,34 +32808,34 @@ l_Lean_Lsp_instToJsonCompletionList___closed__1 = _init_l_Lean_Lsp_instToJsonCom lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList___closed__1); l_Lean_Lsp_instToJsonCompletionList = _init_l_Lean_Lsp_instToJsonCompletionList(); lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionList); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2072____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCompletionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2075____closed__14); l_Lean_Lsp_instFromJsonCompletionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCompletionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCompletionParams___closed__1); l_Lean_Lsp_instFromJsonCompletionParams = _init_l_Lean_Lsp_instFromJsonCompletionParams(); @@ -32846,60 +32846,60 @@ l_Lean_Lsp_instToJsonCompletionParams = _init_l_Lean_Lsp_instToJsonCompletionPar lean_mark_persistent(l_Lean_Lsp_instToJsonCompletionParams); l_Lean_Lsp_Hover_range_x3f___default = _init_l_Lean_Lsp_Hover_range_x3f___default(); lean_mark_persistent(l_Lean_Lsp_Hover_range_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2239____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2242____closed__2); l_Lean_Lsp_instToJsonHover___closed__1 = _init_l_Lean_Lsp_instToJsonHover___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonHover___closed__1); l_Lean_Lsp_instToJsonHover = _init_l_Lean_Lsp_instToJsonHover(); lean_mark_persistent(l_Lean_Lsp_instToJsonHover); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2278____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHover____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2281____closed__13); l_Lean_Lsp_instFromJsonHover___closed__1 = _init_l_Lean_Lsp_instFromJsonHover___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHover___closed__1); l_Lean_Lsp_instFromJsonHover = _init_l_Lean_Lsp_instFromJsonHover(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHover); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2389____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonHoverParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2392____closed__8); l_Lean_Lsp_instFromJsonHoverParams___closed__1 = _init_l_Lean_Lsp_instFromJsonHoverParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonHoverParams___closed__1); l_Lean_Lsp_instFromJsonHoverParams = _init_l_Lean_Lsp_instFromJsonHoverParams(); @@ -32908,22 +32908,22 @@ l_Lean_Lsp_instToJsonHoverParams___closed__1 = _init_l_Lean_Lsp_instToJsonHoverP lean_mark_persistent(l_Lean_Lsp_instToJsonHoverParams___closed__1); l_Lean_Lsp_instToJsonHoverParams = _init_l_Lean_Lsp_instToJsonHoverParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonHoverParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2546____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDeclarationParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2549____closed__8); l_Lean_Lsp_instFromJsonDeclarationParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDeclarationParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDeclarationParams___closed__1); l_Lean_Lsp_instFromJsonDeclarationParams = _init_l_Lean_Lsp_instFromJsonDeclarationParams(); @@ -32932,22 +32932,22 @@ l_Lean_Lsp_instToJsonDeclarationParams___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonDeclarationParams___closed__1); l_Lean_Lsp_instToJsonDeclarationParams = _init_l_Lean_Lsp_instToJsonDeclarationParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonDeclarationParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2703____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2706____closed__8); l_Lean_Lsp_instFromJsonDefinitionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDefinitionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDefinitionParams___closed__1); l_Lean_Lsp_instFromJsonDefinitionParams = _init_l_Lean_Lsp_instFromJsonDefinitionParams(); @@ -32956,22 +32956,22 @@ l_Lean_Lsp_instToJsonDefinitionParams___closed__1 = _init_l_Lean_Lsp_instToJsonD lean_mark_persistent(l_Lean_Lsp_instToJsonDefinitionParams___closed__1); l_Lean_Lsp_instToJsonDefinitionParams = _init_l_Lean_Lsp_instToJsonDefinitionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonDefinitionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2860____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonTypeDefinitionParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_2863____closed__8); l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1 = _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonTypeDefinitionParams___closed__1); l_Lean_Lsp_instFromJsonTypeDefinitionParams = _init_l_Lean_Lsp_instFromJsonTypeDefinitionParams(); @@ -32980,24 +32980,24 @@ l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams___closed__1); l_Lean_Lsp_instToJsonTypeDefinitionParams = _init_l_Lean_Lsp_instToJsonTypeDefinitionParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonTypeDefinitionParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3017____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceContext____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3020____closed__9); l_Lean_Lsp_instFromJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceContext___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceContext___closed__1); l_Lean_Lsp_instFromJsonReferenceContext = _init_l_Lean_Lsp_instFromJsonReferenceContext(); @@ -33006,32 +33006,32 @@ l_Lean_Lsp_instToJsonReferenceContext___closed__1 = _init_l_Lean_Lsp_instToJsonR lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext___closed__1); l_Lean_Lsp_instToJsonReferenceContext = _init_l_Lean_Lsp_instToJsonReferenceContext(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceContext); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3129____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonReferenceParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3132____closed__13); l_Lean_Lsp_instFromJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instFromJsonReferenceParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonReferenceParams___closed__1); l_Lean_Lsp_instFromJsonReferenceParams = _init_l_Lean_Lsp_instFromJsonReferenceParams(); @@ -33040,24 +33040,24 @@ l_Lean_Lsp_instToJsonReferenceParams___closed__1 = _init_l_Lean_Lsp_instToJsonRe lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams___closed__1); l_Lean_Lsp_instToJsonReferenceParams = _init_l_Lean_Lsp_instToJsonReferenceParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonReferenceParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3339____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonWorkspaceSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3342____closed__9); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1 = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonWorkspaceSymbolParams___closed__1); l_Lean_Lsp_instFromJsonWorkspaceSymbolParams = _init_l_Lean_Lsp_instFromJsonWorkspaceSymbolParams(); @@ -33066,22 +33066,22 @@ l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1 = _init_l_Lean_Lsp_instTo lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceSymbolParams___closed__1); l_Lean_Lsp_instToJsonWorkspaceSymbolParams = _init_l_Lean_Lsp_instToJsonWorkspaceSymbolParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonWorkspaceSymbolParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3443____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentHighlightParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3446____closed__8); l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentHighlightParams___closed__1); l_Lean_Lsp_instFromJsonDocumentHighlightParams = _init_l_Lean_Lsp_instFromJsonDocumentHighlightParams(); @@ -33108,18 +33108,18 @@ l_Lean_Lsp_instToJsonDocumentHighlight___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentHighlight___closed__1); l_Lean_Lsp_instToJsonDocumentHighlight = _init_l_Lean_Lsp_instToJsonDocumentHighlight(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentHighlight); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3699____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_3702____closed__6); l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1 = _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonDocumentSymbolParams___closed__1); l_Lean_Lsp_instFromJsonDocumentSymbolParams = _init_l_Lean_Lsp_instFromJsonDocumentSymbolParams(); @@ -33335,66 +33335,66 @@ l_Lean_Lsp_instToJsonSymbolKind___closed__46 = _init_l_Lean_Lsp_instToJsonSymbol lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolKind___closed__46); l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default = _init_l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default(); lean_mark_persistent(l_Lean_Lsp_DocumentSymbolAux_detail_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4324____rarg___closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonDocumentSymbolAux____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4327____rarg___closed__30); l_Lean_Lsp_instToJsonDocumentSymbol___closed__1 = _init_l_Lean_Lsp_instToJsonDocumentSymbol___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonDocumentSymbol___closed__1); l_Lean_Lsp_instToJsonDocumentSymbol = _init_l_Lean_Lsp_instToJsonDocumentSymbol(); @@ -33419,54 +33419,54 @@ l_Lean_Lsp_SymbolInformation_tags___default = _init_l_Lean_Lsp_SymbolInformation lean_mark_persistent(l_Lean_Lsp_SymbolInformation_tags___default); l_Lean_Lsp_SymbolInformation_containerName_x3f___default = _init_l_Lean_Lsp_SymbolInformation_containerName_x3f___default(); lean_mark_persistent(l_Lean_Lsp_SymbolInformation_containerName_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4938____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSymbolInformation____x40_Lean_Data_Lsp_LanguageFeatures___hyg_4941____closed__24); l_Lean_Lsp_instFromJsonSymbolInformation___closed__1 = _init_l_Lean_Lsp_instFromJsonSymbolInformation___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSymbolInformation___closed__1); l_Lean_Lsp_instFromJsonSymbolInformation = _init_l_Lean_Lsp_instFromJsonSymbolInformation(); @@ -33475,22 +33475,22 @@ l_Lean_Lsp_instToJsonSymbolInformation___closed__1 = _init_l_Lean_Lsp_instToJson lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolInformation___closed__1); l_Lean_Lsp_instToJsonSymbolInformation = _init_l_Lean_Lsp_instToJsonSymbolInformation(); lean_mark_persistent(l_Lean_Lsp_instToJsonSymbolInformation); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5247____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyPrepareParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5250____closed__8); l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyPrepareParams(); @@ -33505,58 +33505,58 @@ l_Lean_Lsp_CallHierarchyItem_detail_x3f___default = _init_l_Lean_Lsp_CallHierarc lean_mark_persistent(l_Lean_Lsp_CallHierarchyItem_detail_x3f___default); l_Lean_Lsp_CallHierarchyItem_data_x3f___default = _init_l_Lean_Lsp_CallHierarchyItem_data_x3f___default(); lean_mark_persistent(l_Lean_Lsp_CallHierarchyItem_data_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5469____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_5472____closed__26); l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyItem___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyItem = _init_l_Lean_Lsp_instFromJsonCallHierarchyItem(); @@ -33569,7 +33569,7 @@ l_Lean_Lsp_instBEqCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instBEqCallHi lean_mark_persistent(l_Lean_Lsp_instBEqCallHierarchyItem___closed__1); l_Lean_Lsp_instBEqCallHierarchyItem = _init_l_Lean_Lsp_instBEqCallHierarchyItem(); lean_mark_persistent(l_Lean_Lsp_instBEqCallHierarchyItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6085____closed__1(); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_hashCallHierarchyItem____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6088____closed__1(); l_Lean_Lsp_instHashableCallHierarchyItem___closed__1 = _init_l_Lean_Lsp_instHashableCallHierarchyItem___closed__1(); lean_mark_persistent(l_Lean_Lsp_instHashableCallHierarchyItem___closed__1); l_Lean_Lsp_instHashableCallHierarchyItem = _init_l_Lean_Lsp_instHashableCallHierarchyItem(); @@ -33582,24 +33582,24 @@ l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3 = _init_l_Lean_Lsp_instInh lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyItem___closed__3); l_Lean_Lsp_instInhabitedCallHierarchyItem = _init_l_Lean_Lsp_instInhabitedCallHierarchyItem(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyItem); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6206____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6209____closed__9); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCallsParams(); @@ -33608,34 +33608,34 @@ l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1 = _init_l_Lean lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams___closed__1); l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams = _init_l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyIncomingCallsParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6319____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyIncomingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6322____closed__14); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall = _init_l_Lean_Lsp_instFromJsonCallHierarchyIncomingCall(); @@ -33648,18 +33648,18 @@ l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__1 = _init_l_Lean_Lsp lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall___closed__1); l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall = _init_l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyIncomingCall); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6486____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCallsParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6489____closed__6); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCallsParams(); @@ -33668,28 +33668,28 @@ l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1 = _init_l_Lean lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams___closed__1); l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams = _init_l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonCallHierarchyOutgoingCallsParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6599____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonCallHierarchyOutgoingCall____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6602____closed__11); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1 = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall___closed__1); l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall = _init_l_Lean_Lsp_instFromJsonCallHierarchyOutgoingCall(); @@ -33702,164 +33702,164 @@ l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1 = _init_l_Lean_Lsp lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall___closed__1); l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall = _init_l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall(); lean_mark_persistent(l_Lean_Lsp_instInhabitedCallHierarchyOutgoingCall); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__19); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__20); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__21); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__22); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__23); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__24); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__25); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__26); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__27); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__28); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__29); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__30); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__31); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__32); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__33); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__34); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__35); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__36); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__37); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__38); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__39); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__40); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__41); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__42); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__43); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__44); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__45); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__46); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__47); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6759____closed__48); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__20); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__21); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__22); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__23); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__24); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__25); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__26); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__27); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__28); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__29); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__30); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__31); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__32); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__33); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__34); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__35); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__36); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__37); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__38); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__39); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__40); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__41); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__42); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__43); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__44); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__45); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__46); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__47); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6762____closed__48); l_Lean_Lsp_instToJsonSemanticTokenType___closed__1 = _init_l_Lean_Lsp_instToJsonSemanticTokenType___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenType___closed__1); l_Lean_Lsp_instToJsonSemanticTokenType = _init_l_Lean_Lsp_instToJsonSemanticTokenType(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenType); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__1___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__2___closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__3___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__4___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__5___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__6___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__7___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__8___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__9___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__10___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__11___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__12___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__13___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__14___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__15___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__16___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__17___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__18___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__19___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__20___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__21___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__22___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__23___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____lambda__24___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6918____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__1___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__2___closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__3___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__4___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__5___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__6___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__7___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__8___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__9___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__10___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__11___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__12___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__13___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__14___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__15___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__16___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__17___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__18___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__19___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__20___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__21___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__22___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__23___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____lambda__24___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenType____x40_Lean_Data_Lsp_LanguageFeatures___hyg_6921____closed__2); l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokenType___closed__1); l_Lean_Lsp_instFromJsonSemanticTokenType = _init_l_Lean_Lsp_instFromJsonSemanticTokenType(); @@ -33924,70 +33924,70 @@ l_Lean_Lsp_SemanticTokenType_names___closed__25 = _init_l_Lean_Lsp_SemanticToken lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names___closed__25); l_Lean_Lsp_SemanticTokenType_names = _init_l_Lean_Lsp_SemanticTokenType_names(); lean_mark_persistent(l_Lean_Lsp_SemanticTokenType_names); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__16); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__17); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__18); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7975____closed__19); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__17); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__18); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_7978____closed__19); l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1 = _init_l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenModifier___closed__1); l_Lean_Lsp_instToJsonSemanticTokenModifier = _init_l_Lean_Lsp_instToJsonSemanticTokenModifier(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokenModifier); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__1___closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__2___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__3___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__4___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__5___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__6___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__7___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__8___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____lambda__9___closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8050____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__1___closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__2___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__3___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__4___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__5___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__6___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__7___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__8___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____lambda__9___closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokenModifier____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8053____closed__1); l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokenModifier___closed__1); l_Lean_Lsp_instFromJsonSemanticTokenModifier = _init_l_Lean_Lsp_instFromJsonSemanticTokenModifier(); @@ -34016,34 +34016,34 @@ l_Lean_Lsp_SemanticTokenModifier_names___closed__11 = _init_l_Lean_Lsp_SemanticT lean_mark_persistent(l_Lean_Lsp_SemanticTokenModifier_names___closed__11); l_Lean_Lsp_SemanticTokenModifier_names = _init_l_Lean_Lsp_SemanticTokenModifier_names(); lean_mark_persistent(l_Lean_Lsp_SemanticTokenModifier_names); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8551____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensLegend____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8554____closed__14); l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensLegend = _init_l_Lean_Lsp_instFromJsonSemanticTokensLegend(); @@ -34052,38 +34052,38 @@ l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend___closed__1); l_Lean_Lsp_instToJsonSemanticTokensLegend = _init_l_Lean_Lsp_instToJsonSemanticTokensLegend(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensLegend); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__14); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__15); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8724____closed__16); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__15); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8727____closed__16); l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensOptions___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensOptions = _init_l_Lean_Lsp_instFromJsonSemanticTokensOptions(); @@ -34092,18 +34092,18 @@ l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1 = _init_l_Lean_Lsp_instTo lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensOptions___closed__1); l_Lean_Lsp_instToJsonSemanticTokensOptions = _init_l_Lean_Lsp_instToJsonSemanticTokensOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensOptions); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8934____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_8937____closed__6); l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensParams___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensParams = _init_l_Lean_Lsp_instFromJsonSemanticTokensParams(); @@ -34112,22 +34112,22 @@ l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1 = _init_l_Lean_Lsp_instToJ lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensParams___closed__1); l_Lean_Lsp_instToJsonSemanticTokensParams = _init_l_Lean_Lsp_instToJsonSemanticTokensParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9046____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokensRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9049____closed__8); l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokensRangeParams___closed__1); l_Lean_Lsp_instFromJsonSemanticTokensRangeParams = _init_l_Lean_Lsp_instFromJsonSemanticTokensRangeParams(); @@ -34138,34 +34138,34 @@ l_Lean_Lsp_instToJsonSemanticTokensRangeParams = _init_l_Lean_Lsp_instToJsonSema lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokensRangeParams); l_Lean_Lsp_SemanticTokens_resultId_x3f___default = _init_l_Lean_Lsp_SemanticTokens_resultId_x3f___default(); lean_mark_persistent(l_Lean_Lsp_SemanticTokens_resultId_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__13); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9214____closed__14); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonSemanticTokens____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9217____closed__14); l_Lean_Lsp_instFromJsonSemanticTokens___closed__1 = _init_l_Lean_Lsp_instFromJsonSemanticTokens___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonSemanticTokens___closed__1); l_Lean_Lsp_instFromJsonSemanticTokens = _init_l_Lean_Lsp_instFromJsonSemanticTokens(); @@ -34174,18 +34174,18 @@ l_Lean_Lsp_instToJsonSemanticTokens___closed__1 = _init_l_Lean_Lsp_instToJsonSem lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokens___closed__1); l_Lean_Lsp_instToJsonSemanticTokens = _init_l_Lean_Lsp_instToJsonSemanticTokens(); lean_mark_persistent(l_Lean_Lsp_instToJsonSemanticTokens); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9364____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonFoldingRangeParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9367____closed__6); l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1 = _init_l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonFoldingRangeParams___closed__1); l_Lean_Lsp_instFromJsonFoldingRangeParams = _init_l_Lean_Lsp_instFromJsonFoldingRangeParams(); @@ -34204,33 +34204,33 @@ l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4 = _init_l_Lean_Lsp_instToJsonF lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRangeKind___closed__4); l_Lean_Lsp_FoldingRange_kind_x3f___default = _init_l_Lean_Lsp_FoldingRange_kind_x3f___default(); lean_mark_persistent(l_Lean_Lsp_FoldingRange_kind_x3f___default); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9521____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_toJsonFoldingRange____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9524____closed__2); l_Lean_Lsp_instToJsonFoldingRange___closed__1 = _init_l_Lean_Lsp_instToJsonFoldingRange___closed__1(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRange___closed__1); l_Lean_Lsp_instToJsonFoldingRange = _init_l_Lean_Lsp_instToJsonFoldingRange(); lean_mark_persistent(l_Lean_Lsp_instToJsonFoldingRange); l_Lean_Lsp_RenameOptions_prepareProvider___default = _init_l_Lean_Lsp_RenameOptions_prepareProvider___default(); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9586____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameOptions____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9589____closed__9); l_Lean_Lsp_instFromJsonRenameOptions___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameOptions___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameOptions___closed__1); l_Lean_Lsp_instFromJsonRenameOptions = _init_l_Lean_Lsp_instFromJsonRenameOptions(); @@ -34239,32 +34239,32 @@ l_Lean_Lsp_instToJsonRenameOptions___closed__1 = _init_l_Lean_Lsp_instToJsonRena lean_mark_persistent(l_Lean_Lsp_instToJsonRenameOptions___closed__1); l_Lean_Lsp_instToJsonRenameOptions = _init_l_Lean_Lsp_instToJsonRenameOptions(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameOptions); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__8); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__9); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__10); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__11); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__12); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9698____closed__13); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__9); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__10); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__11); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__12); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9701____closed__13); l_Lean_Lsp_instFromJsonRenameParams___closed__1 = _init_l_Lean_Lsp_instFromJsonRenameParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonRenameParams___closed__1); l_Lean_Lsp_instFromJsonRenameParams = _init_l_Lean_Lsp_instFromJsonRenameParams(); @@ -34273,22 +34273,22 @@ l_Lean_Lsp_instToJsonRenameParams___closed__1 = _init_l_Lean_Lsp_instToJsonRenam lean_mark_persistent(l_Lean_Lsp_instToJsonRenameParams___closed__1); l_Lean_Lsp_instToJsonRenameParams = _init_l_Lean_Lsp_instToJsonRenameParams(); lean_mark_persistent(l_Lean_Lsp_instToJsonRenameParams); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__1); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__2); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__3); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__4); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__5); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__6); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__7); -l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8(); -lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9908____closed__8); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__1); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__2); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__3); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__4); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__5); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__6); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__7); +l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8 = _init_l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8(); +lean_mark_persistent(l___private_Lean_Data_Lsp_LanguageFeatures_0__Lean_Lsp_fromJsonPrepareRenameParams____x40_Lean_Data_Lsp_LanguageFeatures___hyg_9911____closed__8); l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1 = _init_l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1(); lean_mark_persistent(l_Lean_Lsp_instFromJsonPrepareRenameParams___closed__1); l_Lean_Lsp_instFromJsonPrepareRenameParams = _init_l_Lean_Lsp_instFromJsonPrepareRenameParams(); diff --git a/stage0/stdlib/Lean/Declaration.c b/stage0/stdlib/Lean/Declaration.c index a8beb59abd1f..cc327ff6f99d 100644 --- a/stage0/stdlib/Lean/Declaration.c +++ b/stage0/stdlib/Lean/Declaration.c @@ -22,7 +22,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstantVal____ LEAN_EXPORT lean_object* l_Lean_RecursorVal_isUnsafeEx___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592_(uint8_t, lean_object*); LEAN_EXPORT uint8_t lean_inductive_val_is_nested(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkInductiveValEx___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_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_ConstantInfo_value_x21___closed__2; static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__10; LEAN_EXPORT lean_object* l_Lean_instInhabitedRecursorRule; @@ -36,6 +38,7 @@ static lean_object* l_Lean_instInhabitedOpaqueVal___closed__1; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqConstantVal; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkAxiomValEx___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedRecursorVal___closed__1; static lean_object* l_Lean_instReprDefinitionSafety___closed__1; @@ -46,7 +49,6 @@ static lean_object* l_Lean_instInhabitedDeclaration___closed__1; LEAN_EXPORT lean_object* l_Lean_mkRecursorValEx___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_instInhabitedTheoremVal; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x3f___boxed(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991_(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_reducibility_hints_regular(uint32_t); lean_object* l_List_foldlM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_quot_val(lean_object*, lean_object*, lean_object*, uint8_t); @@ -57,6 +59,7 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2_ static lean_object* l_Lean_mkRecName___closed__1; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__4___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedDeclaration; LEAN_EXPORT lean_object* l_Lean_instInhabitedQuotVal; LEAN_EXPORT lean_object* l_Lean_RecursorVal_getFirstMinorIdx(lean_object*); @@ -77,11 +80,11 @@ LEAN_EXPORT lean_object* l_Lean_DefinitionVal_getSafetyEx___boxed(lean_object*); LEAN_EXPORT uint8_t lean_inductive_val_is_unsafe(lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ReducibilityHints_compare(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_toConstantVal___boxed(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_343_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_value_x3f(lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200_(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_inductive_decl(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_instBEqConstantVal___closed__1; LEAN_EXPORT uint8_t lean_quot_val_kind(lean_object*); @@ -90,6 +93,7 @@ LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__3_ LEAN_EXPORT lean_object* l_Lean_RecursorVal_kEx___boxed(lean_object*); static lean_object* l_Lean_instInhabitedConstantVal___closed__1; LEAN_EXPORT lean_object* l_Lean_Declaration_forExprM___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240_(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_definition_val_get_safety(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__20; LEAN_EXPORT lean_object* l_Lean_instBEqDefinitionSafety; @@ -105,6 +109,7 @@ LEAN_EXPORT lean_object* l_Lean_QuotKind_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_RecursorVal_getFirstIndexIdx(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__5; LEAN_EXPORT lean_object* l_Lean_AxiomVal_isUnsafeEx___boxed(lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427_(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__19; LEAN_EXPORT lean_object* l_Lean_InductiveVal_isUnsafeEx___boxed(lean_object*); lean_object* lean_nat_to_int(lean_object*); @@ -117,9 +122,9 @@ lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isCtor(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedConstantVal___closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqDefinitionSafety____x40_Lean_Declaration___hyg_576_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_toCtorIdx(uint8_t); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_QuotKind_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_DefinitionSafety_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_InductiveVal_numCtors___boxed(lean_object*); @@ -135,7 +140,6 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_Declaration_forExprM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTheoremVal___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019_(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_axiom_val_is_unsafe(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40_Lean_Declaration___hyg_465____boxed(lean_object*, lean_object*); LEAN_EXPORT uint32_t lean_reducibility_hints_get_height(lean_object*); @@ -147,12 +151,12 @@ lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__7; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqDefinitionSafety____x40_Lean_Declaration___hyg_576____boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__16; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__4___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkReducibilityHintsRegularEx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_RecursorVal_getFirstIndexIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_isRegular___boxed(lean_object*); @@ -187,7 +191,6 @@ static lean_object* l_Lean_ReducibilityHints_instOrd___closed__1; LEAN_EXPORT uint8_t lean_is_unsafe_inductive_decl(lean_object*); LEAN_EXPORT uint8_t lean_inductive_val_is_reflexive(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200____boxed(lean_object*, lean_object*); static lean_object* l_Lean_ConstantInfo_value_x21___closed__1; static lean_object* l_Lean_instBEqRecursorVal___closed__1; LEAN_EXPORT lean_object* l_Lean_QuotKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); @@ -195,7 +198,6 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__8; static lean_object* l_Lean_instInhabitedInductiveVal___closed__1; LEAN_EXPORT lean_object* l_Lean_instInhabitedConstantInfo; -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387_(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isInductive(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__4(lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Declaration_forExprM___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -203,11 +205,11 @@ uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_name___boxed(lean_object*); static lean_object* l_Lean_DefinitionSafety_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_recursor_val(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__13; LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_isAbbrev___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_ConstantInfo_hasValue___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqConstructorVal; LEAN_EXPORT lean_object* l_Lean_instReprDefinitionSafety; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_hints(lean_object*); @@ -250,7 +252,6 @@ LEAN_EXPORT uint8_t l_Lean_instInhabitedQuotKind; LEAN_EXPORT lean_object* l_Lean_instBEqOpaqueVal; LEAN_EXPORT lean_object* l_Lean_instInhabitedAxiomVal; uint8_t lean_nat_dec_le(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019____boxed(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_ConstantInfo_isPartial(lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__15; @@ -258,7 +259,6 @@ LEAN_EXPORT lean_object* lean_mk_opaque_val(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_ConstantInfo_isPartial___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Lean_QuotKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ReducibilityHints_getHeightEx___boxed(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqAxiomVal____x40_Lean_Declaration___hyg_465_(lean_object*, lean_object*); static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety____x40_Lean_Declaration___hyg_592____closed__2; @@ -267,6 +267,7 @@ static lean_object* l___private_Lean_Declaration_0__Lean_reprDefinitionSafety___ LEAN_EXPORT lean_object* l_Lean_instInhabitedInductiveType; LEAN_EXPORT lean_object* l_Lean_instInhabitedConstantVal; LEAN_EXPORT uint8_t l_Lean_ReducibilityHints_isRegular(lean_object*); +LEAN_EXPORT lean_object* lean_mk_theorem_val(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_inductive_val_is_rec(lean_object*); static lean_object* l_Lean_instInhabitedConstructor___closed__1; LEAN_EXPORT lean_object* l_Lean_ConstantInfo_hints___boxed(lean_object*); @@ -1483,6 +1484,21 @@ x_1 = l_Lean_instBEqTheoremVal___closed__1; return x_1; } } +LEAN_EXPORT lean_object* lean_mk_theorem_val(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +x_7 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_7, 0, x_6); +lean_ctor_set(x_7, 1, x_4); +lean_ctor_set(x_7, 2, x_5); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_OpaqueVal_all___default(lean_object* x_1) { _start: { @@ -1529,7 +1545,7 @@ x_1 = l_Lean_instInhabitedOpaqueVal___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11; @@ -1594,11 +1610,11 @@ return x_18; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -1609,7 +1625,7 @@ static lean_object* _init_l_Lean_instBEqOpaqueVal___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1019____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqOpaqueVal____x40_Lean_Declaration___hyg_1059____boxed), 2, 0); return x_1; } } @@ -2475,7 +2491,7 @@ x_1 = l_Lean_instInhabitedConstructorVal___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; @@ -2566,11 +2582,11 @@ return x_14; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2581,7 +2597,7 @@ static lean_object* _init_l_Lean_instBEqConstructorVal___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_1991____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031____boxed), 2, 0); return x_1; } } @@ -2661,7 +2677,7 @@ x_1 = l_Lean_instInhabitedRecursorRule___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240_(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; @@ -2697,11 +2713,11 @@ return x_13; } } } -LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2712,7 +2728,7 @@ static lean_object* _init_l_Lean_instBEqRecursorRule___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240____boxed), 2, 0); return x_1; } } @@ -2753,7 +2769,7 @@ x_1 = l_Lean_instInhabitedRecursorVal___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____spec__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -2786,7 +2802,7 @@ x_6 = lean_ctor_get(x_1, 0); x_7 = lean_ctor_get(x_1, 1); x_8 = lean_ctor_get(x_2, 0); x_9 = lean_ctor_get(x_2, 1); -x_10 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2200_(x_6, x_8); +x_10 = l___private_Lean_Declaration_0__Lean_beqRecursorRule____x40_Lean_Declaration___hyg_2240_(x_6, x_8); if (x_10 == 0) { uint8_t x_11; @@ -2803,7 +2819,7 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427_(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; 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; uint8_t x_19; uint8_t x_20; lean_object* x_21; uint8_t x_25; @@ -2885,7 +2901,7 @@ return x_36; else { uint8_t x_37; -x_37 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1(x_9, x_18); +x_37 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____spec__1(x_9, x_18); if (x_37 == 0) { uint8_t x_38; @@ -2958,22 +2974,22 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____spec__1___boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____spec__1(x_1, x_2); +x_3 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____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_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387_(x_1, x_2); +x_3 = l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -2984,7 +3000,7 @@ static lean_object* _init_l_Lean_instBEqRecursorVal___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2387____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Declaration_0__Lean_beqRecursorVal____x40_Lean_Declaration___hyg_2427____boxed), 2, 0); return x_1; } } @@ -3610,7 +3626,7 @@ static lean_object* _init_l_Lean_ConstantInfo_value_x21___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_ConstantInfo_value_x21___closed__1; x_2 = l_Lean_ConstantInfo_value_x21___closed__2; -x_3 = lean_unsigned_to_nat(444u); +x_3 = lean_unsigned_to_nat(449u); x_4 = lean_unsigned_to_nat(33u); x_5 = l_Lean_ConstantInfo_value_x21___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3769,7 +3785,7 @@ static lean_object* _init_l_Lean_ConstantInfo_inductiveVal_x21___closed__3() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_ConstantInfo_value_x21___closed__1; x_2 = l_Lean_ConstantInfo_inductiveVal_x21___closed__1; -x_3 = lean_unsigned_to_nat(460u); +x_3 = lean_unsigned_to_nat(465u); x_4 = lean_unsigned_to_nat(9u); x_5 = l_Lean_ConstantInfo_inductiveVal_x21___closed__2; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index 3ece28801aba..721ac54cd93e 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -33,7 +33,6 @@ static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__14; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__3; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor___closed__5; lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__5___boxed(lean_object*, lean_object*); @@ -41,6 +40,7 @@ lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_o static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppLVals___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___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*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_getArgExpectedType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -148,7 +148,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___clos LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_elabAndAddNewArg___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_ElabAppArgs_State_etaArgs___default; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabExplicitUniv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__5; lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_consumeImplicits(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_mkMotive___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,6 +248,7 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__17; 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___regBuiltin_Lean_Elab_Term_elabExplicitUniv_declRange___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Term_ElabElim_finalize___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,7 +282,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn__ static lean_object* l_Lean_Elab_Term_elabAppArgs___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ElabElim_revertArgs___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13; static lean_object* l_panic___at_Lean_Elab_Term_ElabElim_finalize___spec__5___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_revertArgs(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*); @@ -371,11 +370,10 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at___private_Lean_Elab_App_0__ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFn___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__16(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_State_motive_x3f___default; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at_Lean_Elab_Term_ElabElim_revertArgs___spec__6___lambda__1___boxed(lean_object**); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__15; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__15; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithoutExpectedTypeAttr; LEAN_EXPORT uint8_t l_Lean_Elab_Term_elabAppArgs_isFirstOrder___lambda__1(lean_object*); @@ -388,6 +386,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringNamedArg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppFnId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__10; lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern___closed__2; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__18___closed__2; @@ -480,13 +479,11 @@ static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_ lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__12___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice_declRange___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValLoop(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_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ElabElim_revertArgs___spec__11(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__9; lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__2___closed__1; 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*); @@ -512,11 +509,13 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange___c static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_App_0__Lean_Elab_Term_getSuccesses___spec__1___closed__1; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___lambda__2___closed__2; static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__4___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_normalizeFunType(lean_object*); lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj___closed__3; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_elabAppAux___spec__5___closed__1; uint8_t lean_has_out_params(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__5; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processExplictArg___closed__1; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____lambda__2___closed__27; static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__6; @@ -549,7 +548,6 @@ static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__9; LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___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_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ElabElim_revertArgs___spec__7___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* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_ElabElim_finalize___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -649,7 +647,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNe LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_processImplicitArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__4___closed__4; lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____lambda__2___closed__29; static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult_isOutParamOfLocalInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -675,6 +672,7 @@ uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); lean_object* lean_expr_consume_type_annotations(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addEtaArg___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -692,6 +690,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeProj_declRange___close lean_object* l_Lean_Elab_Term_applyResult___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabChoice___closed__3; lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(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_App___hyg_20870____closed__13; static lean_object* l_List_foldlM___at_Lean_Elab_Term_ElabElim_revertArgs___spec__17___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit_declRange___closed__1; @@ -714,6 +713,7 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_throwLValError__ 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___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__2___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__9(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870_(lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__2; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____closed__5; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwInvalidNamedArg___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -721,6 +721,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs lean_object* l_Lean_Elab_Term_postponeElabTerm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_fvarId(lean_object*); static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___lambda__3___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__7; static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__3; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_ElabElim_mkMotive___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -920,7 +921,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabElim_getNextArg_x3f(lean_object*, lean_object* l_Lean_indentExpr(lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__19___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_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__14; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicitUniv___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__2___closed__1; @@ -986,11 +986,13 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNamedPattern_declRange___c lean_object* l_Lean_Expr_getAppFn(lean_object*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at_Lean_Elab_Term_ElabElim_revertArgs___spec__6(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__4; LEAN_EXPORT lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_finalize___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_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__3; uint8_t l_Lean_Expr_isProp(lean_object*); static lean_object* l_Lean_Elab_Term_elabExplicit___closed__2; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____lambda__2___closed__31; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__14; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____lambda__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ElabElim_revertArgs___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_EXPORT lean_object* l_Lean_Elab_Term_elabProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1022,6 +1024,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_findBinderN LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ensureArgType(lean_object*, 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_App___hyg_6707____closed__4; uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_typeMatchesBaseName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__8; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__3___closed__2; @@ -1040,6 +1043,7 @@ static lean_object* l_Lean_Elab_Term_ElabElim_finalize___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Term_ElabElim_mkMotive___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_Elab_Term_eraseNamedArg___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__9; static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_propagateExpectedType___spec__1___closed__1; static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_throwInvalidFieldNotation___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs_elabAsElim_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1052,6 +1056,7 @@ static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_observing___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabProj_declRange___closed__5; @@ -1117,13 +1122,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_App_0__Lean_ lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_shouldElabAsElim(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAsElim; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__7; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabApp_declRange___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_normalizeFunType___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_App_0__Lean_Elab_Term_hasOptAutoParams___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_initFn____x40_Lean_Elab_App___hyg_20864____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_addLValArg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_6707____lambda__2___closed__22; @@ -1145,7 +1147,6 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_elabAppArgs___spe static lean_object* l_List_forIn_loop___at___private_Lean_Elab_App_0__Lean_Elab_Term_mkBaseProjections___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_ElabElim_revertArgs___spec__16___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_addNewArg(lean_object*, lean_object*, 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_App___hyg_20864____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_ElabElim_mkMotive___spec__1___rarg___lambda__1(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___private_Lean_Elab_App_0__Lean_Elab_Term_resolveDotName_go___closed__4; @@ -1169,6 +1170,7 @@ lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext LEAN_EXPORT uint8_t l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_shouldPropagateExpectedTypeFor(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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__12; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ElabAppArgs_trySynthesizeAppInstMVars___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_ElabElim_finalize___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ElabAppArgs_synthesizeAppInstMVars___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1182,7 +1184,6 @@ uint8_t l_Lean_Expr_isConst(lean_object*); 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___private_Lean_Elab_App_0__Lean_Elab_Term_resolveLValAux___lambda__5___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult_isOutParamOfLocalInstance___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_throwInvalidNamedArg___rarg___closed__2; static lean_object* l_Lean_Elab_Term_elabAppArgs___lambda__5___closed__2; @@ -1219,7 +1220,6 @@ static lean_object* l___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNe LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_isNextOutParamOfLocalInstanceAndResult___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_App_0__Lean_Elab_Term_ElabAppArgs_anyNamedArgDependsOnCurrent___spec__14(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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabExplicit___closed__2; LEAN_EXPORT lean_object* l_Lean_getRefPos___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__5___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_App_0__Lean_Elab_Term_mergeFailures___spec__2___closed__1; @@ -50130,7 +50130,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -50140,7 +50140,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -50150,27 +50150,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_App___hyg_20864____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____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_App___hyg_20864____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__5() { _start: { lean_object* x_1; @@ -50178,17 +50178,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_App___hyg_20864____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__5; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__7() { _start: { lean_object* x_1; @@ -50196,37 +50196,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__9; x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_5____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11() { _start: { lean_object* x_1; @@ -50234,17 +50234,17 @@ x_1 = lean_mk_string_from_bytes("App", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__13() { _start: { lean_object* x_1; @@ -50252,33 +50252,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_App___hyg_20864____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____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_App___hyg_20864____closed__14; -x_2 = lean_unsigned_to_nat(20864u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__14; +x_2 = lean_unsigned_to_nat(20870u); 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_App___hyg_20864_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870_(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_App___hyg_20864____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -51238,37 +51238,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabProj_declRange___closed__ if (builtin) {res = l___regBuiltin_Lean_Elab_Term_elabProj_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20864_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_App___hyg_20870_(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/Binders.c b/stage0/stdlib/Lean/Elab/Binders.c index 0ba8cfd1e04f..81d6dca81ce4 100644 --- a/stage0/stdlib/Lean/Elab/Binders.c +++ b/stage0/stdlib/Lean/Elab/Binders.c @@ -15,6 +15,7 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabFun___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__2; @@ -22,19 +23,16 @@ lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabFun___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__5; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__38; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabArrow___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_expandFunBinders_loop___spec__1(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__13; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux_loop(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Syntax_mkNameLit(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_expandForall___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__2; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__1; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__33; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_precheckFun___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -46,7 +44,7 @@ static lean_object* l_Lean_Elab_Term_expandLetEqnsDecl___closed__3; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_elabLetDeclCore___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___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__5; lean_object* l_Lean_addBuiltinDocString(lean_object*, lean_object*, lean_object*); @@ -59,9 +57,9 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__7; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__6; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDepArrow(lean_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_elabLetDeclCore(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -76,9 +74,11 @@ static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__7; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__2; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11; +lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__2; -lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,22 +88,26 @@ static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange(lean_object*); +static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__1; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__52; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_elabFunBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__3; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___closed__1; uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16; +lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__5; lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___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*); @@ -114,15 +118,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBin LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_expandFunBinders_loop___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_expandForall___closed__7; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__12; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__24; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandFun(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instMonadQuotationOfMonadFunctorOfMonadLift___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__40; lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); lean_object* l_Lean_compileDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFunBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -138,7 +143,6 @@ static lean_object* l_Lean_Elab_Term_elabArrow___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -151,9 +155,10 @@ LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetDeclAux___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_quoteAutoTactic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_quoteAutoTactic(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__4; lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,7 +166,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabLetDec lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__3(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__32; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__2; @@ -170,11 +174,12 @@ static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__1; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__11; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandForall(lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckFun(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*); @@ -186,6 +191,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___boxed(lea LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderType___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__1; static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__3; +lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFun___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -213,11 +220,9 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__2; -static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4; lean_object* l_Lean_Syntax_getNumArgs(lean_object*); lean_object* l_Lean_Elab_Term_mkExplicitBinder(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3; uint8_t l_Lean_Name_isImplementationDetail(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabFun___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandLetEqnsDecl(lean_object*, uint8_t, lean_object*, lean_object*); @@ -230,15 +235,15 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderView LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__2; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__7; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandFunBinders_loop___spec__2(lean_object*, lean_object*); @@ -251,7 +256,6 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_Quotation_withNewLocals___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1(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_Term_quoteAutoTactic___closed__39; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__3; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_elabFun___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDeclsOpt(lean_object*, lean_object*, lean_object*, lean_object*); @@ -260,6 +264,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabFun___spec__3 static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__3; +lean_object* l_Lean_mkStrLit(lean_object*); lean_object* l_Lean_Macro_resolveGlobalName(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__4; @@ -289,24 +294,22 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_elabLetDecl lean_object* l_List_forM___at_Lean_Elab_Term_Quotation_precheck___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__6; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__54; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_ReaderT_instMonadFunctor(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_unzip___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux_loop___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__45; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__21; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandFunBinders_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; lean_object* l_Lean_mkHole(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__1; lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__22; @@ -317,14 +320,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux(lean_object*, lean_obje lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__3; -lean_object* l_Lean_Syntax_mkStrLit(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__47; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__4; lean_object* l_Lean_Elab_Term_expandOptType(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__48; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___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_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__3; @@ -333,14 +335,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_elabFunBindersAux___boxed(l LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); -lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__3; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__43; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__49; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckFun___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDepArrow___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -349,17 +348,15 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_elabLetDecl static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___rarg___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun___closed__2; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967_(lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10; lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabFun___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_quoteNameMk(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkBinderAnnotations; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___closed__2; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForall___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__2; @@ -385,8 +382,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__1; extern lean_object* l_Lean_Elab_Term_termElabAttribute; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6; +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckFun___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*); @@ -404,6 +400,8 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinder static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandForall___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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_Binders_0__Lean_Elab_Term_getBinderIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -412,9 +410,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatchTactic___boxed static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__3; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__30; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__7; -lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkLetIdDeclView(lean_object*); uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__3; @@ -422,6 +418,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_docString(le LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__7___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkLetIdDeclView___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFunBinders___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__29; @@ -431,6 +428,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatch(lean_object*, LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun(lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_kindOfBinderName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___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*); @@ -438,15 +437,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls(lean_object* LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_expandWhereDecls___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__4; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__11; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__46; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9; extern lean_object* l_Lean_Elab_macroAttribute; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__36; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); @@ -456,14 +456,18 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_ 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*); uint8_t l_Lean_Syntax_isIdent(lean_object*); lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabFun___spec__5___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__1; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckFun___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__4; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__7; @@ -472,28 +476,23 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderView LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__2; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___boxed(lean_object**); static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__2; +static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckFun___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_Term_expandFunBinders_loop___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__3; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3; -static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__1; -extern lean_object* l_Lean_Core_instMonadQuotationCoreM; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow___closed__3; @@ -506,47 +505,47 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderView LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1(lean_object*); static lean_object* l_Lean_Elab_Term_expandForall___closed__3; lean_object* l_Lean_Elab_Term_Quotation_precheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_instToExprPreresolved___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__3; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__2; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5; LEAN_EXPORT lean_object* l_Array_sequenceMap___at_Lean_Elab_Term_expandWhereDecls___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__4; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__50; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3; lean_object* l_Lean_Syntax_setArg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl(lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14; lean_object* l_Lean_Elab_Term_addTermInfo_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__1; lean_object* l_Lean_Meta_mkLetFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__4(lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckFun___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_propagateExpectedType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_universeConstraintsCheckpoint___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__37; static lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addLocalVarInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__25; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__41; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__7(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_State_expectedType_x3f___default; -static lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_propagateExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandFunBinders_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__6; uint8_t l_Lean_Name_isAtomic(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5___rarg(lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__35; +lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckFun___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*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___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*); @@ -555,26 +554,31 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__ LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_State_fvars___default; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__5(lean_object*, size_t, size_t, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__34; +lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__2; lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_SepArray_ofElems(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5; lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10; lean_object* lean_environment_main_module(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___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___regBuiltin_Lean_Elab_Term_precheckFun___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForall(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandForall___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__7; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__6; lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -582,11 +586,9 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange__ static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___closed__2; -lean_object* l_StateRefT_x27_lift(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_Elab_Term_elabArrow___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -595,7 +597,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__13; 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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_kindOfBinderName___boxed(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; @@ -606,6 +607,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___close LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__4; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__12; @@ -619,43 +621,44 @@ lean_object* l_Lean_Elab_Term_clearInMatch(lean_object*, lean_object*, lean_obje static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__26; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__31; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__19; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__1; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__3; +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFunBinders_loop___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Term_expandForall___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___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___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__4; -lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_ToExpr_0__Lean_Name_toExprAux(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15; -lean_object* l_String_intercalate(lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isClass_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__51; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDelayedDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___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_Term_elabLetDeclAux___lambda__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__9; static lean_object* l_Lean_Elab_Term_expandLetEqnsDecl___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFun___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__10; +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunBinderIds_x3f(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___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*); @@ -669,11 +672,11 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderVi static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__8; @@ -684,8 +687,8 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRan static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__4; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__21; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderType(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__53; static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__6; lean_object* l_Lean_Elab_Term_getMatchAltsNumPatterns(lean_object*); @@ -694,7 +697,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___close LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__2; -lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__3; @@ -706,13 +708,13 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__6; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9; LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__2___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_elabLetDeclCore___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_elabFun___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -723,12 +725,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinders___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__2; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -739,7 +739,6 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinder LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinders(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339_(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMatchAltsWhereDecls_loop___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__6; @@ -750,6 +749,7 @@ lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_Quotation_pr LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDecls(lean_object*, 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_Elab_Term_expandForall___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); @@ -757,7 +757,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed_ lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__15; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__44; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__9; @@ -767,27 +766,25 @@ static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandExplicitFun(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed__2; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__1; -static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__42; lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isAntiquotSplice(lean_object*); static lean_object* l_Lean_Elab_Term_declareTacticSyntax___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__7; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5; lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_expandFunBinders_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5; lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__5; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabFun(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__2; uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderType(lean_object* x_1, lean_object* x_2) { _start: { @@ -1266,88 +1263,197 @@ x_3 = lean_box(x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(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("app", 3); -return x_1; -} +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(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); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2() { -_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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; -x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; -x_4 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +lean_inc(x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3() { +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3(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("Array.push", 10); -return x_1; -} +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(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); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3; -x_2 = l_String_toSubstring_x27(x_1); -return x_2; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +lean_inc(x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5() { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Array", 5); -return x_1; +uint8_t x_6; +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 5); +x_8 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_dec(x_1); +lean_ctor_set(x_3, 5, x_8); +x_9 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3(x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_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; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_3, 2); +x_13 = lean_ctor_get(x_3, 3); +x_14 = lean_ctor_get(x_3, 4); +x_15 = lean_ctor_get(x_3, 5); +x_16 = lean_ctor_get(x_3, 6); +x_17 = lean_ctor_get(x_3, 7); +x_18 = lean_ctor_get(x_3, 8); +x_19 = lean_ctor_get(x_3, 9); +x_20 = lean_ctor_get(x_3, 10); +x_21 = lean_ctor_get_uint8(x_3, sizeof(void*)*11); +x_22 = lean_ctor_get_uint8(x_3, sizeof(void*)*11 + 1); +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_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +x_23 = l_Lean_replaceRef(x_1, x_15); +lean_dec(x_15); +lean_dec(x_1); +x_24 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_24, 0, x_10); +lean_ctor_set(x_24, 1, x_11); +lean_ctor_set(x_24, 2, x_12); +lean_ctor_set(x_24, 3, x_13); +lean_ctor_set(x_24, 4, x_14); +lean_ctor_set(x_24, 5, x_23); +lean_ctor_set(x_24, 6, x_16); +lean_ctor_set(x_24, 7, x_17); +lean_ctor_set(x_24, 8, x_18); +lean_ctor_set(x_24, 9, x_19); +lean_ctor_set(x_24, 10, x_20); +lean_ctor_set_uint8(x_24, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_24, sizeof(void*)*11 + 1, x_22); +x_25 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3(x_2, x_24, x_4, x_5); +lean_dec(x_4); +lean_dec(x_24); +return x_25; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6() { +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("push", 4); +x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); +x_1 = lean_box(0); +x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("null", 4); +x_1 = lean_mk_string_from_bytes("Array", 5); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("push", 4); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6() { _start: { lean_object* x_1; @@ -1355,1340 +1461,1013 @@ x_1 = lean_mk_string_from_bytes("invalid auto tactic, antiquotation is not allow return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, size_t x_9, size_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_19; -x_19 = lean_usize_dec_lt(x_10, x_9); -if (x_19 == 0) +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_6, x_5); +if (x_11 == 0) { -lean_object* x_20; -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_7); -lean_dec(x_6); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_11); -lean_ctor_set(x_20, 1, x_18); -return x_20; +lean_object* x_12; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_7); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_53; lean_object* x_61; uint8_t x_62; -x_21 = lean_array_uget(x_8, x_10); -x_61 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -x_62 = lean_name_eq(x_1, x_61); -if (x_62 == 0) +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_21; uint8_t x_22; +x_13 = lean_array_uget(x_4, x_6); +x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; +x_22 = lean_name_eq(x_1, x_21); +if (x_22 == 0) { -lean_object* x_63; -x_63 = lean_box(0); -x_22 = x_63; -goto block_52; +lean_object* x_23; +lean_inc(x_9); +lean_inc(x_8); +x_23 = l_Lean_Elab_Term_quoteAutoTactic(x_13, x_8, x_9, x_10); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5; +lean_inc(x_3); +x_27 = l_Lean_Expr_const___override(x_26, x_3); +lean_inc(x_2); +x_28 = l_Lean_mkApp3(x_27, x_2, x_7, x_24); +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_28); +x_14 = x_29; +x_15 = x_25; +goto block_20; } else { -uint8_t x_64; -x_64 = l_Lean_Syntax_isAntiquotSuffixSplice(x_21); -if (x_64 == 0) -{ -uint8_t x_65; -x_65 = l_Lean_Syntax_isAntiquotSplice(x_21); -if (x_65 == 0) +uint8_t x_30; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_30 = !lean_is_exclusive(x_23); +if (x_30 == 0) { -lean_object* x_66; -x_66 = lean_box(0); -x_22 = x_66; -goto block_52; +return x_23; } else { -lean_object* x_67; -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -x_67 = lean_box(0); -x_53 = x_67; -goto block_60; +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; +} } } else { -lean_object* x_68; -lean_dec(x_11); -lean_dec(x_7); -lean_dec(x_6); -x_68 = lean_box(0); -x_53 = x_68; -goto block_60; -} +uint8_t x_34; +x_34 = l_Lean_Syntax_isAntiquotSuffixSplice(x_13); +if (x_34 == 0) +{ +uint8_t x_35; +x_35 = l_Lean_Syntax_isAntiquotSplice(x_13); +if (x_35 == 0) +{ +lean_object* x_36; +lean_inc(x_9); +lean_inc(x_8); +x_36 = l_Lean_Elab_Term_quoteAutoTactic(x_13, x_8, x_9, x_10); +if (lean_obj_tag(x_36) == 0) +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +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_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5; +lean_inc(x_3); +x_40 = l_Lean_Expr_const___override(x_39, x_3); +lean_inc(x_2); +x_41 = l_Lean_mkApp3(x_40, x_2, x_7, x_37); +x_42 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_14 = x_42; +x_15 = x_38; +goto block_20; } -block_52: +else { -lean_object* x_23; -lean_dec(x_22); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); -x_23 = l_Lean_Elab_Term_quoteAutoTactic(x_21, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_23) == 0) +uint8_t x_43; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_43 = !lean_is_exclusive(x_36); +if (x_43 == 0) { -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; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; size_t x_45; size_t x_46; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_ctor_get(x_16, 5); -lean_inc(x_26); -x_27 = 0; -x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); -x_29 = lean_ctor_get(x_16, 10); -lean_inc(x_29); -x_30 = lean_st_ref_get(x_17, x_25); -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_ctor_get(x_31, 0); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_environment_main_module(x_33); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7; -x_36 = l_Lean_addMacroScope(x_34, x_35, x_29); -lean_inc(x_6); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_6); -lean_inc(x_7); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_7); -x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4; -lean_inc(x_28); -x_40 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_40, 0, x_28); -lean_ctor_set(x_40, 1, x_39); -lean_ctor_set(x_40, 2, x_36); -lean_ctor_set(x_40, 3, x_38); -x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_28); -x_42 = l_Lean_Syntax_node2(x_28, x_41, x_11, x_24); -x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_44 = l_Lean_Syntax_node2(x_28, x_43, x_40, x_42); -x_45 = 1; -x_46 = lean_usize_add(x_10, x_45); -x_10 = x_46; -x_11 = x_44; -x_18 = x_32; -goto _start; +return x_36; } else { -uint8_t x_48; -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_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_36, 0); +x_45 = lean_ctor_get(x_36, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_36); +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 +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_dec(x_7); -lean_dec(x_6); -x_48 = !lean_is_exclusive(x_23); -if (x_48 == 0) +lean_dec(x_3); +lean_dec(x_2); +x_47 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7; +x_48 = l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_13, x_47, x_8, x_9, x_10); +x_49 = !lean_is_exclusive(x_48); +if (x_49 == 0) { -return x_23; +return x_48; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_23, 0); -x_50 = lean_ctor_get(x_23, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_48, 0); +x_51 = lean_ctor_get(x_48, 1); +lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_23); -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_dec(x_48); +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; } } } -block_60: +else { -lean_object* x_54; lean_object* x_55; uint8_t x_56; -lean_dec(x_53); -x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11; -x_55 = l_Lean_throwErrorAt___at_Lean_Elab_Term_Quotation_getAntiquotationIds___spec__1(x_21, x_54, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_17); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_21); -x_56 = !lean_is_exclusive(x_55); -if (x_56 == 0) +lean_object* x_53; lean_object* x_54; uint8_t x_55; +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_53 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7; +x_54 = l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__2(x_13, x_53, x_8, x_9, x_10); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -return x_55; +return x_54; } 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_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_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_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; +} } } +block_20: +{ +lean_object* x_16; size_t x_17; size_t x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_6, x_17); +x_6 = x_18; +x_7 = x_16; +x_10 = x_15; +goto _start; +} } } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__1() { +LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid auto tactic, tactic is missing", 38); -return x_1; -} +uint8_t x_6; +x_6 = !lean_is_exclusive(x_3); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_3, 5); +x_8 = l_Lean_replaceRef(x_1, x_7); +lean_dec(x_7); +lean_dec(x_1); +lean_ctor_set(x_3, 5, x_8); +x_9 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +x_12 = lean_ctor_get(x_3, 2); +x_13 = lean_ctor_get(x_3, 3); +x_14 = lean_ctor_get(x_3, 4); +x_15 = lean_ctor_get(x_3, 5); +x_16 = lean_ctor_get(x_3, 6); +x_17 = lean_ctor_get(x_3, 7); +x_18 = lean_ctor_get(x_3, 8); +x_19 = lean_ctor_get(x_3, 9); +x_20 = lean_ctor_get(x_3, 10); +x_21 = lean_ctor_get_uint8(x_3, sizeof(void*)*11); +x_22 = lean_ctor_get_uint8(x_3, sizeof(void*)*11 + 1); +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_12); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_3); +x_23 = l_Lean_replaceRef(x_1, x_15); +lean_dec(x_15); +lean_dec(x_1); +x_24 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_24, 0, x_10); +lean_ctor_set(x_24, 1, x_11); +lean_ctor_set(x_24, 2, x_12); +lean_ctor_set(x_24, 3, x_13); +lean_ctor_set(x_24, 4, x_14); +lean_ctor_set(x_24, 5, x_23); +lean_ctor_set(x_24, 6, x_16); +lean_ctor_set(x_24, 7, x_17); +lean_ctor_set(x_24, 8, x_18); +lean_ctor_set(x_24, 9, x_19); +lean_ctor_set(x_24, 10, x_20); +lean_ctor_set_uint8(x_24, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_24, sizeof(void*)*11 + 1, x_22); +x_25 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_2, x_24, x_4, x_5); +lean_dec(x_4); +lean_dec(x_24); +return x_25; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__3() { +} +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_ReaderT_instMonadFunctor), 3, 2); -lean_closure_set(x_1, 0, lean_box(0)); -lean_closure_set(x_1, 1, lean_box(0)); +x_1 = lean_mk_string_from_bytes("Syntax", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__4() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_ReaderT_instMonadLift), 3, 2); -lean_closure_set(x_1, 0, lean_box(0)); -lean_closure_set(x_1, 1, lean_box(0)); +x_1 = lean_mk_string_from_bytes("Preresolved", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__5() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_StateRefT_x27_lift), 4, 3); -lean_closure_set(x_1, 0, lean_box(0)); -lean_closure_set(x_1, 1, lean_box(0)); -lean_closure_set(x_1, 2, lean_box(0)); +x_1 = lean_mk_string_from_bytes("namespace", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__6() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__5; -x_3 = l_Lean_Core_instMonadQuotationCoreM; -x_4 = l_Lean_instMonadQuotationOfMonadFunctorOfMonadLift___rarg(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2; +x_4 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__7() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__6; -x_4 = l_Lean_instMonadQuotationOfMonadFunctorOfMonadLift___rarg(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__8() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__5; -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__7; -x_4 = l_Lean_instMonadQuotationOfMonadFunctorOfMonadLift___rarg(x_1, x_2, x_3); -return x_4; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("decl", 4); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__9() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2; +x_4 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__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_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__7; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term", 4); +x_1 = lean_mk_string_from_bytes("String", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__10() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__9; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__11() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__10; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__12() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Array.empty", 11); +x_1 = lean_mk_string_from_bytes("List", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__12; -x_2 = l_String_toSubstring_x27(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__14() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("empty", 5); +x_1 = lean_mk_string_from_bytes("nil", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__15() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__14; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__16() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; +x_2 = l_Lean_levelZero; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; -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_quoteAutoTactic___closed__18() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15; +x_3 = l_Lean_Expr_const___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__19() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__18; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11; +x_3 = l_Lean_Expr_app___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__20() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Syntax.node", 11); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; -x_2 = l_String_toSubstring_x27(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__22() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Syntax", 6); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__23() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("node", 4); +x_1 = lean_mk_string_from_bytes("cons", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__24() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__25() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15; +x_3 = l_Lean_Expr_const___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__26() { +static lean_object* _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___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_Term_quoteAutoTactic___closed__25; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11; +x_3 = l_Lean_Expr_app___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__27() { +LEAN_EXPORT lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__25; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +if (lean_obj_tag(x_3) == 0) +{ +lean_dec(x_2); +lean_inc(x_1); +return x_1; } +else +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +lean_inc(x_2); +x_6 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6(x_1, x_2, x_5); +if (lean_obj_tag(x_4) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +lean_dec(x_4); +x_8 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_7); +x_9 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5; +x_10 = l_Lean_Expr_app___override(x_9, x_8); +x_11 = l_Lean_mkAppB(x_2, x_10, x_6); +return x_11; } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__28() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_12 = lean_ctor_get(x_4, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_4, 1); +lean_inc(x_13); +lean_dec(x_4); +x_14 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_12); +x_15 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17; +x_16 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__21; +x_17 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_instToExprPreresolved___spec__1(x_15, x_16, x_13); +x_18 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8; +x_19 = l_Lean_mkAppB(x_18, x_14, x_17); +x_20 = l_Lean_mkAppB(x_2, x_19, x_6); +return x_20; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__29() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__28; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__30() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("SourceInfo.none", 15); +x_1 = lean_mk_string_from_bytes("invalid auto tactic, tactic is missing", 38); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__31() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__30; -x_2 = l_String_toSubstring_x27(x_1); +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__32() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__3() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("SourceInfo", 10); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__33() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__5() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("none", 4); +x_1 = lean_mk_string_from_bytes("empty", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__34() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__32; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__5; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__35() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__32; -x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__33; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__36() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__35; +x_2 = lean_box(0); x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__37() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___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_Term_quoteAutoTactic___closed__38() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___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); +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__6; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__7; +x_3 = l_Lean_Expr_const___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__39() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__36; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___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); +x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__8; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; +x_3 = l_Lean_Expr_app___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__40() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__10() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("quotedName", 10); +x_1 = lean_mk_string_from_bytes("node", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__41() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__11() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; -x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; -x_4 = l_Lean_Elab_Term_quoteAutoTactic___closed__40; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__10; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__42() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__12() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(".", 1); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__11; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__43() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__13() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`", 1); +x_1 = lean_mk_string_from_bytes("SourceInfo", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__44() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__45() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__14() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("mkAtom", 6); +x_1 = lean_mk_string_from_bytes("none", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__46() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__15() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; -x_2 = l_String_toSubstring_x27(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__13; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__14; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__47() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; -x_3 = l_Lean_Name_str___override(x_1, x_2); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__48() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("mkAtom", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__45; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__17; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__49() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_quoteAutoTactic___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); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__18; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__50() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__20() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__48; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__51() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___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_Term_quoteAutoTactic___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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__52() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__49; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___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); +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; +x_3 = l_Lean_Expr_const___override(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__53() { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__23() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("invalid auto tactic, identifier is not allowed", 46); +x_1 = lean_mk_string_from_bytes("toSubstring", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__54() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_quoteAutoTactic___closed__53; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_quoteAutoTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__24() { _start: { -switch (lean_obj_tag(x_1)) { -case 0: -{ -lean_object* x_9; lean_object* x_10; -x_9 = l_Lean_Elab_Term_quoteAutoTactic___closed__2; -x_10 = l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__11(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; -} -case 1: -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 2); -lean_inc(x_12); -x_13 = l_Lean_Syntax_isAntiquot(x_1); -if (x_13 == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_1); -if (x_14 == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; 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; size_t x_34; size_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_15 = lean_ctor_get(x_1, 2); -lean_dec(x_15); -x_16 = lean_ctor_get(x_1, 1); -lean_dec(x_16); -x_17 = lean_ctor_get(x_1, 0); -lean_dec(x_17); -x_18 = lean_ctor_get(x_6, 5); -lean_inc(x_18); -x_19 = 0; -x_20 = l_Lean_SourceInfo_fromRef(x_18, x_19); -x_21 = lean_ctor_get(x_6, 10); -lean_inc(x_21); -x_22 = lean_st_ref_get(x_7, x_8); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); -lean_dec(x_22); -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_environment_main_module(x_25); -x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; -lean_inc(x_21); -x_29 = l_Lean_addMacroScope(x_26, x_28, x_21); -x_30 = l_Lean_Elab_Term_quoteAutoTactic___closed__13; -x_31 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; -lean_inc(x_20); -x_32 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_32, 0, x_20); -lean_ctor_set(x_32, 1, x_30); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_32, 3, x_31); -x_33 = lean_array_get_size(x_12); -x_34 = lean_usize_of_nat(x_33); -lean_dec(x_33); -x_35 = 0; -x_36 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; -x_37 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; -x_38 = l_Lean_Elab_Term_quoteAutoTactic___closed__8; -x_39 = l_Lean_Elab_Term_quoteAutoTactic___closed__11; -lean_inc(x_7); -x_40 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_11, x_36, x_37, x_38, x_39, x_27, x_27, x_12, x_34, x_35, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_24); -lean_dec(x_12); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_st_ref_get(x_7, x_42); -lean_dec(x_7); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec(x_45); -x_47 = lean_environment_main_module(x_46); -x_48 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; -lean_inc(x_21); -lean_inc(x_47); -x_49 = l_Lean_addMacroScope(x_47, x_48, x_21); -x_50 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; -x_51 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; -lean_inc(x_20); -x_52 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_52, 0, x_20); -lean_ctor_set(x_52, 1, x_50); -lean_ctor_set(x_52, 2, x_49); -lean_ctor_set(x_52, 3, x_51); -x_53 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; -x_54 = l_Lean_addMacroScope(x_47, x_53, x_21); -x_55 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; -x_56 = l_Lean_Elab_Term_quoteAutoTactic___closed__39; -lean_inc(x_20); -x_57 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_57, 0, x_20); -lean_ctor_set(x_57, 1, x_55); -lean_ctor_set(x_57, 2, x_54); -lean_ctor_set(x_57, 3, x_56); -lean_inc(x_11); -x_58 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_27, x_11); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -lean_free_object(x_1); -x_59 = l_Lean_quoteNameMk(x_11); -x_60 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_20); -x_61 = l_Lean_Syntax_node3(x_20, x_60, x_57, x_59, x_41); -x_62 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_63 = l_Lean_Syntax_node2(x_20, x_62, x_52, x_61); -lean_ctor_set(x_43, 0, x_63); -return x_43; -} -else -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_11); -x_64 = lean_ctor_get(x_58, 0); -lean_inc(x_64); -lean_dec(x_58); -x_65 = l_Lean_Elab_Term_quoteAutoTactic___closed__42; -x_66 = l_String_intercalate(x_65, x_64); -x_67 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; -x_68 = lean_string_append(x_67, x_66); -lean_dec(x_66); -x_69 = lean_box(2); -x_70 = l_Lean_Syntax_mkNameLit(x_68, x_69); -x_71 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; -x_72 = lean_array_push(x_71, x_70); -x_73 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; -lean_ctor_set(x_1, 2, x_72); -lean_ctor_set(x_1, 1, x_73); -lean_ctor_set(x_1, 0, x_69); -x_74 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_20); -x_75 = l_Lean_Syntax_node3(x_20, x_74, x_57, x_1, x_41); -x_76 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_77 = l_Lean_Syntax_node2(x_20, x_76, x_52, x_75); -lean_ctor_set(x_43, 0, x_77); -return x_43; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__23; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -else -{ -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_78 = lean_ctor_get(x_43, 0); -x_79 = lean_ctor_get(x_43, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_43); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -lean_dec(x_78); -x_81 = lean_environment_main_module(x_80); -x_82 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; -lean_inc(x_21); -lean_inc(x_81); -x_83 = l_Lean_addMacroScope(x_81, x_82, x_21); -x_84 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; -x_85 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; -lean_inc(x_20); -x_86 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_86, 0, x_20); -lean_ctor_set(x_86, 1, x_84); -lean_ctor_set(x_86, 2, x_83); -lean_ctor_set(x_86, 3, x_85); -x_87 = l_Lean_Elab_Term_quoteAutoTactic___closed__34; -x_88 = l_Lean_addMacroScope(x_81, x_87, x_21); -x_89 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; -x_90 = l_Lean_Elab_Term_quoteAutoTactic___closed__39; -lean_inc(x_20); -x_91 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_91, 0, x_20); -lean_ctor_set(x_91, 1, x_89); -lean_ctor_set(x_91, 2, x_88); -lean_ctor_set(x_91, 3, x_90); -lean_inc(x_11); -x_92 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_27, x_11); -if (lean_obj_tag(x_92) == 0) +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__25() { +_start: { -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_free_object(x_1); -x_93 = l_Lean_quoteNameMk(x_11); -x_94 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_20); -x_95 = l_Lean_Syntax_node3(x_20, x_94, x_91, x_93, x_41); -x_96 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_97 = l_Lean_Syntax_node2(x_20, x_96, x_86, x_95); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_79); -return x_98; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } -else -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -lean_dec(x_11); -x_99 = lean_ctor_get(x_92, 0); -lean_inc(x_99); -lean_dec(x_92); -x_100 = l_Lean_Elab_Term_quoteAutoTactic___closed__42; -x_101 = l_String_intercalate(x_100, x_99); -x_102 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; -x_103 = lean_string_append(x_102, x_101); -lean_dec(x_101); -x_104 = lean_box(2); -x_105 = l_Lean_Syntax_mkNameLit(x_103, x_104); -x_106 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; -x_107 = lean_array_push(x_106, x_105); -x_108 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; -lean_ctor_set(x_1, 2, x_107); -lean_ctor_set(x_1, 1, x_108); -lean_ctor_set(x_1, 0, x_104); -x_109 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_20); -x_110 = l_Lean_Syntax_node3(x_20, x_109, x_91, x_1, x_41); -x_111 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_112 = l_Lean_Syntax_node2(x_20, x_111, x_86, x_110); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_79); -return x_113; } +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__26() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1; +x_3 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; } } -else +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__27() { +_start: { -uint8_t x_114; -lean_dec(x_21); -lean_dec(x_20); -lean_free_object(x_1); -lean_dec(x_11); -lean_dec(x_7); -x_114 = !lean_is_exclusive(x_40); -if (x_114 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__26; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__28() { +_start: { -return x_40; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; +x_3 = l_Lean_Expr_app___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Elab_Term_quoteAutoTactic___closed__29() { +_start: { -lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_115 = lean_ctor_get(x_40, 0); -x_116 = lean_ctor_get(x_40, 1); -lean_inc(x_116); -lean_inc(x_115); -lean_dec(x_40); -x_117 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -return x_117; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__27; +x_3 = l_Lean_Expr_app___override(x_1, x_2); +return x_3; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_quoteAutoTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_Elab_Term_quoteAutoTactic___closed__2; +x_6 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_5, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_6; } -else +case 1: { -lean_object* x_118; uint8_t 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; size_t x_134; size_t x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_1); -x_118 = lean_ctor_get(x_6, 5); -lean_inc(x_118); -x_119 = 0; -x_120 = l_Lean_SourceInfo_fromRef(x_118, x_119); -x_121 = lean_ctor_get(x_6, 10); -lean_inc(x_121); -x_122 = lean_st_ref_get(x_7, x_8); -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_ctor_get(x_123, 0); -lean_inc(x_125); -lean_dec(x_123); -x_126 = lean_environment_main_module(x_125); -x_127 = lean_box(0); -x_128 = l_Lean_Elab_Term_quoteAutoTactic___closed__15; -lean_inc(x_121); -x_129 = l_Lean_addMacroScope(x_126, x_128, x_121); -x_130 = l_Lean_Elab_Term_quoteAutoTactic___closed__13; -x_131 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; -lean_inc(x_120); -x_132 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_132, 0, x_120); -lean_ctor_set(x_132, 1, x_130); -lean_ctor_set(x_132, 2, x_129); -lean_ctor_set(x_132, 3, x_131); -x_133 = lean_array_get_size(x_12); -x_134 = lean_usize_of_nat(x_133); -lean_dec(x_133); -x_135 = 0; -x_136 = l_Lean_Elab_Term_quoteAutoTactic___closed__3; -x_137 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; -x_138 = l_Lean_Elab_Term_quoteAutoTactic___closed__8; -x_139 = l_Lean_Elab_Term_quoteAutoTactic___closed__11; +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = lean_ctor_get(x_1, 1); lean_inc(x_7); -x_140 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_11, x_136, x_137, x_138, x_139, x_127, x_127, x_12, x_134, x_135, x_132, x_2, x_3, x_4, x_5, x_6, x_7, x_124); -lean_dec(x_12); -if (lean_obj_tag(x_140) == 0) +x_8 = lean_ctor_get(x_1, 2); +lean_inc(x_8); +x_9 = l_Lean_Syntax_isAntiquot(x_1); +if (x_9 == 0) { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; -x_141 = lean_ctor_get(x_140, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_140, 1); -lean_inc(x_142); -lean_dec(x_140); -x_143 = lean_st_ref_get(x_7, x_142); -lean_dec(x_7); -x_144 = lean_ctor_get(x_143, 0); -lean_inc(x_144); -x_145 = lean_ctor_get(x_143, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_143)) { - lean_ctor_release(x_143, 0); - lean_ctor_release(x_143, 1); - x_146 = x_143; -} else { - lean_dec_ref(x_143); - x_146 = lean_box(0); -} -x_147 = lean_ctor_get(x_144, 0); -lean_inc(x_147); -lean_dec(x_144); -x_148 = lean_environment_main_module(x_147); -x_149 = l_Lean_Elab_Term_quoteAutoTactic___closed__24; -lean_inc(x_121); -lean_inc(x_148); -x_150 = l_Lean_addMacroScope(x_148, x_149, x_121); -x_151 = l_Lean_Elab_Term_quoteAutoTactic___closed__21; -x_152 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; -lean_inc(x_120); -x_153 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_153, 0, x_120); -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_Term_quoteAutoTactic___closed__34; -x_155 = l_Lean_addMacroScope(x_148, x_154, x_121); -x_156 = l_Lean_Elab_Term_quoteAutoTactic___closed__31; -x_157 = l_Lean_Elab_Term_quoteAutoTactic___closed__39; -lean_inc(x_120); -x_158 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_158, 0, x_120); -lean_ctor_set(x_158, 1, x_156); -lean_ctor_set(x_158, 2, x_155); -lean_ctor_set(x_158, 3, x_157); -lean_inc(x_11); -x_159 = l___private_Init_Meta_0__Lean_getEscapedNameParts_x3f(x_127, x_11); -if (lean_obj_tag(x_159) == 0) -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_160 = l_Lean_quoteNameMk(x_11); -x_161 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_120); -x_162 = l_Lean_Syntax_node3(x_120, x_161, x_158, x_160, x_141); -x_163 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_164 = l_Lean_Syntax_node2(x_120, x_163, x_153, x_162); -if (lean_is_scalar(x_146)) { - x_165 = lean_alloc_ctor(0, 2, 0); -} else { - x_165 = x_146; -} -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_145); -return x_165; +lean_object* x_10; size_t x_11; size_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_dec(x_1); +x_10 = lean_array_get_size(x_8); +x_11 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_12 = 0; +x_13 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; +x_14 = l_Lean_Elab_Term_quoteAutoTactic___closed__7; +x_15 = l_Lean_Elab_Term_quoteAutoTactic___closed__9; +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4(x_7, x_13, x_14, x_8, x_11, x_12, x_15, x_2, x_3, x_4); +lean_dec(x_8); +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_16, 0); +x_19 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_7); +x_20 = l_Lean_Elab_Term_quoteAutoTactic___closed__12; +x_21 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; +x_22 = l_Lean_mkApp3(x_20, x_21, x_19, x_18); +lean_ctor_set(x_16, 0, x_22); +return x_16; } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; 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_dec(x_11); -x_166 = lean_ctor_get(x_159, 0); -lean_inc(x_166); -lean_dec(x_159); -x_167 = l_Lean_Elab_Term_quoteAutoTactic___closed__42; -x_168 = l_String_intercalate(x_167, x_166); -x_169 = l_Lean_Elab_Term_quoteAutoTactic___closed__43; -x_170 = lean_string_append(x_169, x_168); -lean_dec(x_168); -x_171 = lean_box(2); -x_172 = l_Lean_Syntax_mkNameLit(x_170, x_171); -x_173 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; -x_174 = lean_array_push(x_173, x_172); -x_175 = l_Lean_Elab_Term_quoteAutoTactic___closed__41; -x_176 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_176, 0, x_171); -lean_ctor_set(x_176, 1, x_175); -lean_ctor_set(x_176, 2, x_174); -x_177 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_120); -x_178 = l_Lean_Syntax_node3(x_120, x_177, x_158, x_176, x_141); -x_179 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_180 = l_Lean_Syntax_node2(x_120, x_179, x_153, x_178); -if (lean_is_scalar(x_146)) { - x_181 = lean_alloc_ctor(0, 2, 0); -} else { - x_181 = x_146; -} -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_145); -return x_181; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_23 = lean_ctor_get(x_16, 0); +x_24 = lean_ctor_get(x_16, 1); +lean_inc(x_24); +lean_inc(x_23); +lean_dec(x_16); +x_25 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_7); +x_26 = l_Lean_Elab_Term_quoteAutoTactic___closed__12; +x_27 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; +x_28 = l_Lean_mkApp3(x_26, x_27, x_25, x_23); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_24); +return x_29; } } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_121); -lean_dec(x_120); -lean_dec(x_11); +uint8_t x_30; lean_dec(x_7); -x_182 = lean_ctor_get(x_140, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_140, 1); -lean_inc(x_183); -if (lean_is_exclusive(x_140)) { - lean_ctor_release(x_140, 0); - lean_ctor_release(x_140, 1); - x_184 = x_140; -} else { - lean_dec_ref(x_140); - x_184 = lean_box(0); -} -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); -} else { - x_185 = x_184; +x_30 = !lean_is_exclusive(x_16); +if (x_30 == 0) +{ +return x_16; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -return x_185; +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_16, 0); +x_32 = lean_ctor_get(x_16, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_16); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } else { -lean_object* x_186; lean_object* x_187; -lean_dec(x_12); -lean_dec(x_11); -x_186 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11; -x_187 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_186, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_187; +lean_object* x_34; lean_object* x_35; +lean_dec(x_8); +lean_dec(x_7); +x_34 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7; +x_35 = l_Lean_throwErrorAt___at_Lean_Elab_Term_quoteAutoTactic___spec__5(x_1, x_34, x_2, x_3, x_4); +return x_35; } } case 2: { -lean_object* x_188; lean_object* x_189; uint8_t x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; uint8_t x_194; -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_dec(x_3); lean_dec(x_2); -x_188 = lean_ctor_get(x_1, 1); -lean_inc(x_188); +x_36 = lean_ctor_get(x_1, 1); +lean_inc(x_36); lean_dec(x_1); -x_189 = lean_ctor_get(x_6, 5); -lean_inc(x_189); -x_190 = 0; -x_191 = l_Lean_SourceInfo_fromRef(x_189, x_190); -x_192 = lean_ctor_get(x_6, 10); -lean_inc(x_192); -lean_dec(x_6); -x_193 = lean_st_ref_get(x_7, x_8); -lean_dec(x_7); -x_194 = !lean_is_exclusive(x_193); -if (x_194 == 0) +x_37 = l_Lean_mkStrLit(x_36); +x_38 = l_Lean_Elab_Term_quoteAutoTactic___closed__19; +x_39 = l_Lean_Expr_app___override(x_38, x_37); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_4); +return x_40; +} +default: { -lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; -x_195 = lean_ctor_get(x_193, 0); -x_196 = lean_ctor_get(x_195, 0); -lean_inc(x_196); -lean_dec(x_195); -x_197 = lean_environment_main_module(x_196); -x_198 = l_Lean_Elab_Term_quoteAutoTactic___closed__47; -x_199 = l_Lean_addMacroScope(x_197, x_198, x_192); -x_200 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; -x_201 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; -lean_inc(x_191); -x_202 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_202, 0, x_191); -lean_ctor_set(x_202, 1, x_200); -lean_ctor_set(x_202, 2, x_199); -lean_ctor_set(x_202, 3, x_201); -x_203 = lean_box(2); -x_204 = l_Lean_Syntax_mkStrLit(x_188, x_203); -lean_dec(x_188); -x_205 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_191); -x_206 = l_Lean_Syntax_node1(x_191, x_205, x_204); -x_207 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_208 = l_Lean_Syntax_node2(x_191, x_207, x_202, x_206); -lean_ctor_set(x_193, 0, x_208); -return x_193; -} -else -{ -lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; -x_209 = lean_ctor_get(x_193, 0); -x_210 = lean_ctor_get(x_193, 1); -lean_inc(x_210); -lean_inc(x_209); -lean_dec(x_193); -x_211 = lean_ctor_get(x_209, 0); -lean_inc(x_211); -lean_dec(x_209); -x_212 = lean_environment_main_module(x_211); -x_213 = l_Lean_Elab_Term_quoteAutoTactic___closed__47; -x_214 = l_Lean_addMacroScope(x_212, x_213, x_192); -x_215 = l_Lean_Elab_Term_quoteAutoTactic___closed__46; -x_216 = l_Lean_Elab_Term_quoteAutoTactic___closed__52; -lean_inc(x_191); -x_217 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_217, 0, x_191); -lean_ctor_set(x_217, 1, x_215); -lean_ctor_set(x_217, 2, x_214); -lean_ctor_set(x_217, 3, x_216); -x_218 = lean_box(2); -x_219 = l_Lean_Syntax_mkStrLit(x_188, x_218); -lean_dec(x_188); -x_220 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; -lean_inc(x_191); -x_221 = l_Lean_Syntax_node1(x_191, x_220, x_219); -x_222 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; -x_223 = l_Lean_Syntax_node2(x_191, x_222, x_217, x_221); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_210); -return x_224; +lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_dec(x_3); +lean_dec(x_2); +x_41 = lean_ctor_get(x_1, 2); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 3); +lean_inc(x_42); +lean_dec(x_1); +x_43 = 1; +lean_inc(x_41); +x_44 = l_Lean_Name_toString(x_41, x_43); +x_45 = l_Lean_mkStrLit(x_44); +x_46 = l_Lean_Elab_Term_quoteAutoTactic___closed__25; +x_47 = l_Lean_Expr_app___override(x_46, x_45); +x_48 = l___private_Lean_ToExpr_0__Lean_Name_toExprAux(x_41); +x_49 = l_Lean_Elab_Term_quoteAutoTactic___closed__28; +x_50 = l_Lean_Elab_Term_quoteAutoTactic___closed__29; +x_51 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6(x_49, x_50, x_42); +x_52 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; +x_53 = l_Lean_Elab_Term_quoteAutoTactic___closed__16; +x_54 = l_Lean_mkApp4(x_52, x_53, x_47, x_48, x_51); +x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_4); +return x_55; } } -default: +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_225; lean_object* x_226; -x_225 = l_Lean_Elab_Term_quoteAutoTactic___closed__54; -x_226 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_1, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_226; +lean_object* x_5; +x_5 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; } } +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwError___at_Lean_Elab_Term_quoteAutoTactic___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___boxed(lean_object** _args) { -lean_object* x_1 = _args[0]; -lean_object* x_2 = _args[1]; -lean_object* x_3 = _args[2]; -lean_object* x_4 = _args[3]; -lean_object* x_5 = _args[4]; -lean_object* x_6 = _args[5]; -lean_object* x_7 = _args[6]; -lean_object* x_8 = _args[7]; -lean_object* x_9 = _args[8]; -lean_object* x_10 = _args[9]; -lean_object* x_11 = _args[10]; -lean_object* x_12 = _args[11]; -lean_object* x_13 = _args[12]; -lean_object* x_14 = _args[13]; -lean_object* x_15 = _args[14]; -lean_object* x_16 = _args[15]; -lean_object* x_17 = _args[16]; -lean_object* x_18 = _args[17]; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -size_t x_19; size_t x_20; lean_object* x_21; -x_19 = lean_unbox_usize(x_9); -lean_dec(x_9); -x_20 = lean_unbox_usize(x_10); -lean_dec(x_10); -x_21 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19, x_20, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -lean_dec(x_8); +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_5); lean_dec(x_5); +x_12 = lean_unbox_usize(x_6); +lean_dec(x_6); +x_13 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4(x_1, x_2, x_3, x_4, x_11, x_12, x_7, x_8, x_9, x_10); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); lean_dec(x_1); -return x_21; +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -2784,55 +2563,25 @@ lean_dec(x_1); x_30 = !lean_is_exclusive(x_19); if (x_30 == 0) { -return x_19; -} -else -{ -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_19, 0); -x_32 = lean_ctor_get(x_19, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_19); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; -} -} -} -} -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__22; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; -} +return x_19; } -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; -x_2 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_19); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; +} +} } } -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -2840,7 +2589,7 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -2848,12 +2597,12 @@ x_1 = lean_mk_string_from_bytes("autoParam", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5; +x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } @@ -2878,120 +2627,63 @@ x_16 = l_Lean_addMacroScope(x_14, x_1, x_15); x_17 = lean_box(0); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_18 = l_Lean_Elab_Term_quoteAutoTactic(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_18 = l_Lean_Elab_Term_quoteAutoTactic(x_2, x_7, x_8, x_12); if (lean_obj_tag(x_18) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +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_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_Term_declareTacticSyntax___lambda__2___closed__3; -x_22 = 1; -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_23 = l_Lean_Elab_Term_elabTerm(x_19, x_21, x_22, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_20); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); +x_22 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_unbox(x_23); lean_dec(x_23); -x_26 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6; -x_30 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_28); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_unbox(x_31); -lean_dec(x_31); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_30, 1); -lean_inc(x_33); -lean_dec(x_30); -x_34 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; -x_35 = lean_box(0); -x_36 = l_Lean_Elab_Term_declareTacticSyntax___lambda__1(x_16, x_17, x_34, x_27, x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_33); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_36; -} -else +if (x_24 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_37 = lean_ctor_get(x_30, 1); -lean_inc(x_37); -lean_dec(x_30); -lean_inc(x_27); -x_38 = l_Lean_MessageData_ofExpr(x_27); -x_39 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_29, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_37); -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 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; -x_43 = l_Lean_Elab_Term_declareTacticSyntax___lambda__1(x_16, x_17, x_42, x_27, x_40, x_3, x_4, x_5, x_6, x_7, x_8, x_41); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_22, 1); +lean_inc(x_25); +lean_dec(x_22); +x_26 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; +x_27 = lean_box(0); +x_28 = l_Lean_Elab_Term_declareTacticSyntax___lambda__1(x_16, x_17, x_26, x_19, x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_25); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_40); -return x_43; -} +return x_28; } else { -uint8_t x_44; -lean_dec(x_16); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_dec(x_22); +lean_inc(x_19); +x_30 = l_Lean_MessageData_ofExpr(x_19); +x_31 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_21, x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_29); +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_Elab_Term_quoteAutoTactic___closed__4; +x_35 = l_Lean_Elab_Term_declareTacticSyntax___lambda__1(x_16, x_17, x_34, x_19, x_32, x_3, x_4, x_5, x_6, x_7, x_8, x_33); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_44 = !lean_is_exclusive(x_23); -if (x_44 == 0) -{ -return x_23; -} -else -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_23, 0); -x_46 = lean_ctor_get(x_23, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_23); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; -} +lean_dec(x_32); +return x_35; } } else { -uint8_t x_48; +uint8_t x_36; lean_dec(x_16); lean_dec(x_8); lean_dec(x_7); @@ -2999,23 +2691,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_48 = !lean_is_exclusive(x_18); -if (x_48 == 0) +x_36 = !lean_is_exclusive(x_18); +if (x_36 == 0) { return x_18; } else { -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_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_18, 0); +x_38 = lean_ctor_get(x_18, 1); +lean_inc(x_38); +lean_inc(x_37); 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; +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; } } } @@ -3110,47 +2802,67 @@ return x_5; static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__5() { _start: { +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("app", 3); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__2; +x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; +x_4 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__5; +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_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7() { +_start: +{ lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5; +x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11() { _start: { lean_object* x_1; @@ -3158,43 +2870,43 @@ x_1 = lean_mk_string_from_bytes("optParam", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9; +x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; x_2 = l_String_toSubstring_x27(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__9; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13() { +static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12; +x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); @@ -3271,10 +2983,10 @@ x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); lean_dec(x_30); x_32 = lean_environment_main_module(x_31); -x_33 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; +x_33 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; x_34 = l_Lean_addMacroScope(x_32, x_33, x_27); -x_35 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__5; -x_36 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; +x_35 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7; +x_36 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10; lean_inc(x_26); x_37 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_37, 0, x_26); @@ -3282,10 +2994,10 @@ lean_ctor_set(x_37, 1, x_35); lean_ctor_set(x_37, 2, x_34); lean_ctor_set(x_37, 3, x_36); x_38 = l_Lean_mkIdentFrom(x_20, x_22, x_25); -x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_26); x_40 = l_Lean_Syntax_node2(x_26, x_39, x_1, x_38); -x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; +x_41 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; x_42 = l_Lean_Syntax_node2(x_26, x_41, x_37, x_40); lean_ctor_set(x_28, 0, x_42); return x_28; @@ -3302,10 +3014,10 @@ x_45 = lean_ctor_get(x_43, 0); lean_inc(x_45); lean_dec(x_43); x_46 = lean_environment_main_module(x_45); -x_47 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; +x_47 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; x_48 = l_Lean_addMacroScope(x_46, x_47, x_27); -x_49 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__5; -x_50 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__8; +x_49 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__7; +x_50 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10; lean_inc(x_26); x_51 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_51, 0, x_26); @@ -3313,10 +3025,10 @@ 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_mkIdentFrom(x_20, x_22, x_25); -x_53 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_53 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_26); x_54 = l_Lean_Syntax_node2(x_26, x_53, x_1, x_52); -x_55 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; +x_55 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; x_56 = l_Lean_Syntax_node2(x_26, x_55, x_51, x_54); x_57 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_57, 0, x_56); @@ -3381,20 +3093,20 @@ x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); lean_dec(x_70); x_72 = lean_environment_main_module(x_71); -x_73 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; +x_73 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; x_74 = l_Lean_addMacroScope(x_72, x_73, x_67); -x_75 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10; -x_76 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; +x_75 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12; +x_76 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15; lean_inc(x_66); x_77 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_77, 0, x_66); lean_ctor_set(x_77, 1, x_75); lean_ctor_set(x_77, 2, x_74); lean_ctor_set(x_77, 3, x_76); -x_78 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_78 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_66); x_79 = l_Lean_Syntax_node2(x_66, x_78, x_1, x_63); -x_80 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; +x_80 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; x_81 = l_Lean_Syntax_node2(x_66, x_80, x_77, x_79); lean_ctor_set(x_68, 0, x_81); return x_68; @@ -3411,20 +3123,20 @@ x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); lean_dec(x_82); x_85 = lean_environment_main_module(x_84); -x_86 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; +x_86 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; x_87 = l_Lean_addMacroScope(x_85, x_86, x_67); -x_88 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__10; -x_89 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13; +x_88 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12; +x_89 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15; lean_inc(x_66); x_90 = lean_alloc_ctor(3, 4, 0); lean_ctor_set(x_90, 0, x_66); 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_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_91 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_66); x_92 = l_Lean_Syntax_node2(x_66, x_91, x_1, x_63); -x_93 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; +x_93 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; x_94 = l_Lean_Syntax_node2(x_66, x_93, x_90, x_92); x_95 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_95, 0, x_94); @@ -3461,22 +3173,14 @@ return x_10; static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ident", 5); -return x_1; -} -} -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; +x_2 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2() { _start: { lean_object* x_1; @@ -3484,11 +3188,11 @@ x_1 = lean_mk_string_from_bytes("identifier or `_` expected", 26); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3; +x_1 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -3520,7 +3224,7 @@ x_14 = lean_unsigned_to_nat(0u); x_15 = lean_array_uset(x_3, x_2, x_14); lean_inc(x_13); x_16 = l_Lean_Syntax_getKind(x_13); -x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; +x_17 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; x_18 = lean_name_eq(x_16, x_17); if (x_18 == 0) { @@ -3532,7 +3236,7 @@ if (x_20 == 0) { lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_dec(x_15); -x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4; +x_21 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3; x_22 = l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(x_13, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) @@ -4012,6 +3716,15 @@ 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_Binders_0__Lean_Elab_Term_toBinderViews___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews(lean_object* x_1, lean_object* x_2, 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: { @@ -4086,7 +3799,7 @@ lean_ctor_set(x_30, 0, x_26); lean_ctor_set(x_30, 1, x_26); lean_ctor_set(x_30, 2, x_28); lean_ctor_set_uint8(x_30, sizeof(void*)*3, x_29); -x_31 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_31 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_32 = lean_array_push(x_31, x_30); lean_ctor_set(x_24, 0, x_32); return x_24; @@ -4109,7 +3822,7 @@ lean_ctor_set(x_38, 0, x_33); lean_ctor_set(x_38, 1, x_33); lean_ctor_set(x_38, 2, x_36); lean_ctor_set_uint8(x_38, sizeof(void*)*3, x_37); -x_39 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_39 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_40 = lean_array_push(x_39, x_38); x_41 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_41, 0, x_40); @@ -4369,7 +4082,7 @@ lean_ctor_set(x_99, 0, x_1); lean_ctor_set(x_99, 1, x_95); lean_ctor_set(x_99, 2, x_97); lean_ctor_set_uint8(x_99, sizeof(void*)*3, x_98); -x_100 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_100 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_101 = lean_array_push(x_100, x_99); lean_ctor_set(x_93, 0, x_101); return x_93; @@ -4391,7 +4104,7 @@ lean_ctor_set(x_107, 0, x_1); lean_ctor_set(x_107, 1, x_102); lean_ctor_set(x_107, 2, x_105); lean_ctor_set_uint8(x_107, sizeof(void*)*3, x_106); -x_108 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_108 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_109 = lean_array_push(x_108, x_107); x_110 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_110, 0, x_109); @@ -4451,7 +4164,7 @@ lean_ctor_set(x_121, 0, x_1); lean_ctor_set(x_121, 1, x_117); lean_ctor_set(x_121, 2, x_119); lean_ctor_set_uint8(x_121, sizeof(void*)*3, x_120); -x_122 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_122 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_123 = lean_array_push(x_122, x_121); lean_ctor_set(x_115, 0, x_123); return x_115; @@ -4473,7 +4186,7 @@ lean_ctor_set(x_129, 0, x_1); lean_ctor_set(x_129, 1, x_124); lean_ctor_set(x_129, 2, x_127); lean_ctor_set_uint8(x_129, sizeof(void*)*3, x_128); -x_130 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_130 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_131 = lean_array_push(x_130, x_129); x_132 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_132, 0, x_131); @@ -4794,7 +4507,7 @@ lean_dec(x_1); return x_10; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1() { _start: { lean_object* x_1; @@ -4802,17 +4515,17 @@ x_1 = lean_mk_string_from_bytes("checkBinderAnnotations", 22); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3() { _start: { lean_object* x_1; @@ -4820,7 +4533,7 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4() { _start: { lean_object* x_1; @@ -4828,13 +4541,13 @@ x_1 = lean_mk_string_from_bytes("check whether type is a class instance whenever return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -4843,25 +4556,25 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2; -x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2; +x_3 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -5961,7 +5674,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder___rarg(lean_object* x_1, le _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_10 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_11 = lean_array_push(x_10, x_1); x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabBinder___rarg___lambda__1___boxed), 9, 1); lean_closure_set(x_12, 0, x_2); @@ -6051,7 +5764,7 @@ lean_inc(x_12); x_14 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_12); x_16 = l_Lean_Syntax_node1(x_12, x_15, x_2); x_17 = l_Lean_Elab_Term_expandSimpleBinderWithType___closed__3; @@ -6093,7 +5806,7 @@ lean_inc(x_29); x_31 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); -x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_29); x_33 = l_Lean_Syntax_node1(x_29, x_32, x_2); x_34 = l_Lean_Elab_Term_expandSimpleBinderWithType___closed__3; @@ -6235,6 +5948,36 @@ static lean_object* _init_l_Lean_Elab_Term_expandForall___closed__5() { _start: { lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandForall___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_expandForall___closed__5; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandForall___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_expandForall___closed__6; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_2); +lean_ctor_set(x_3, 1, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_expandForall___closed__8() { +_start: +{ +lean_object* x_1; x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } @@ -6315,7 +6058,7 @@ x_25 = lean_array_get_size(x_24); x_26 = lean_usize_of_nat(x_25); lean_dec(x_25); x_27 = 0; -x_28 = l_Lean_Elab_Term_quoteAutoTactic___closed__11; +x_28 = l_Lean_Elab_Term_expandForall___closed__7; lean_inc(x_2); x_29 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandForall___spec__1(x_28, x_21, x_26, x_27, x_24, x_2, x_3); if (lean_obj_tag(x_29) == 0) @@ -6338,7 +6081,7 @@ lean_ctor_set(x_36, 0, x_34); lean_ctor_set(x_36, 1, x_35); x_37 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_38 = l_Array_append___rarg(x_37, x_31); -x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_34); x_40 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_40, 0, x_34); @@ -6349,7 +6092,7 @@ x_41 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_41, 0, x_34); lean_ctor_set(x_41, 1, x_39); lean_ctor_set(x_41, 2, x_37); -x_42 = l_Lean_Elab_Term_expandForall___closed__5; +x_42 = l_Lean_Elab_Term_expandForall___closed__8; lean_inc(x_34); x_43 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_43, 0, x_34); @@ -6378,7 +6121,7 @@ lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); x_52 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_53 = l_Array_append___rarg(x_52, x_45); -x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_54 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_49); x_55 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_55, 0, x_49); @@ -6389,7 +6132,7 @@ x_56 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_56, 0, x_49); lean_ctor_set(x_56, 1, x_54); lean_ctor_set(x_56, 2, x_52); -x_57 = l_Lean_Elab_Term_expandForall___closed__5; +x_57 = l_Lean_Elab_Term_expandForall___closed__8; lean_inc(x_49); x_58 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_58, 0, x_49); @@ -6456,7 +6199,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_expandForall___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -6495,7 +6238,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(259u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6507,7 +6250,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(267u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6535,7 +6278,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(259u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6547,7 +6290,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(259u); +x_1 = lean_unsigned_to_nat(263u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6763,7 +6506,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabForall___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -6802,7 +6545,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +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); @@ -6814,7 +6557,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(271u); +x_1 = lean_unsigned_to_nat(275u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6842,7 +6585,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6854,7 +6597,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(265u); +x_1 = lean_unsigned_to_nat(269u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7012,7 +6755,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__ { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -7245,7 +6988,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabArrow___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -7276,7 +7019,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(280u); +x_1 = lean_unsigned_to_nat(284u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7288,7 +7031,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(287u); +x_1 = lean_unsigned_to_nat(291u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7316,7 +7059,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(280u); +x_1 = lean_unsigned_to_nat(284u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7328,7 +7071,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(280u); +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); @@ -7382,7 +7125,7 @@ x_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(2u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); -x_14 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_14 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_15 = lean_array_push(x_14, x_11); x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabForall___lambda__1), 9, 1); lean_closure_set(x_16, 0, x_13); @@ -7433,7 +7176,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__4 { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -7482,7 +7225,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7494,7 +7237,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(298u); +x_1 = lean_unsigned_to_nat(302u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7522,7 +7265,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7534,7 +7277,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(293u); +x_1 = lean_unsigned_to_nat(297u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7665,7 +7408,7 @@ x_12 = l_Lean_Syntax_isOfKind(x_11, x_1); if (x_12 == 0) { lean_object* x_13; uint8_t x_14; -x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; lean_inc(x_11); x_14 = l_Lean_Syntax_isOfKind(x_11, x_13); if (x_14 == 0) @@ -7707,7 +7450,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_getFunB _start: { lean_object* x_4; uint8_t x_5; -x_4 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2; +x_4 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__6; lean_inc(x_1); x_5 = l_Lean_Syntax_isOfKind(x_1, x_4); if (x_5 == 0) @@ -7719,7 +7462,7 @@ x_7 = l_Lean_Syntax_isOfKind(x_1, x_6); if (x_7 == 0) { lean_object* x_8; uint8_t x_9; -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; lean_inc(x_1); x_9 = l_Lean_Syntax_isOfKind(x_1, x_8); if (x_9 == 0) @@ -7774,7 +7517,7 @@ x_26 = l_Lean_Syntax_isOfKind(x_21, x_25); if (x_26 == 0) { lean_object* x_27; uint8_t x_28; -x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2; +x_27 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; lean_inc(x_21); x_28 = l_Lean_Syntax_isOfKind(x_21, x_27); if (x_28 == 0) @@ -8456,7 +8199,7 @@ lean_object* x_17; lean_object* x_18; uint8_t x_19; x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1; +x_18 = l_Lean_Elab_Term_quoteAutoTactic___closed__20; x_19 = lean_string_dec_eq(x_17, x_18); lean_dec(x_17); if (x_19 == 0) @@ -8516,7 +8259,7 @@ lean_inc(x_40); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_43 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_44 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_40); x_45 = lean_alloc_ctor(1, 3, 0); @@ -8580,7 +8323,7 @@ lean_inc(x_68); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_68); lean_ctor_set(x_70, 1, x_69); -x_71 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_71 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_72 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_68); x_73 = lean_alloc_ctor(1, 3, 0); @@ -8657,7 +8400,7 @@ lean_inc(x_99); x_101 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_101, 0, x_99); lean_ctor_set(x_101, 1, x_100); -x_102 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_102 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_103 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_99); x_104 = lean_alloc_ctor(1, 3, 0); @@ -8751,7 +8494,7 @@ lean_inc(x_133); x_135 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_135, 0, x_133); lean_ctor_set(x_135, 1, x_134); -x_136 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_136 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_137 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_133); x_138 = lean_alloc_ctor(1, 3, 0); @@ -8947,7 +8690,7 @@ lean_inc(x_199); x_201 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_201, 0, x_199); lean_ctor_set(x_201, 1, x_200); -x_202 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_202 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_203 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_199); x_204 = lean_alloc_ctor(1, 3, 0); @@ -9011,7 +8754,7 @@ lean_inc(x_227); x_229 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_229, 0, x_227); lean_ctor_set(x_229, 1, x_228); -x_230 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_230 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_231 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_227); x_232 = lean_alloc_ctor(1, 3, 0); @@ -9088,7 +8831,7 @@ lean_inc(x_258); x_260 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_260, 0, x_258); lean_ctor_set(x_260, 1, x_259); -x_261 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_261 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_262 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_258); x_263 = lean_alloc_ctor(1, 3, 0); @@ -9182,7 +8925,7 @@ lean_inc(x_292); x_294 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_294, 0, x_292); lean_ctor_set(x_294, 1, x_293); -x_295 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_295 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_296 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_292); x_297 = lean_alloc_ctor(1, 3, 0); @@ -9338,7 +9081,7 @@ lean_inc(x_347); x_349 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_349, 0, x_347); lean_ctor_set(x_349, 1, x_348); -x_350 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_350 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_351 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_347); x_352 = lean_alloc_ctor(1, 3, 0); @@ -9402,7 +9145,7 @@ lean_inc(x_375); x_377 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_377, 0, x_375); lean_ctor_set(x_377, 1, x_376); -x_378 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_378 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_379 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_375); x_380 = lean_alloc_ctor(1, 3, 0); @@ -9479,7 +9222,7 @@ lean_inc(x_406); x_408 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_408, 0, x_406); lean_ctor_set(x_408, 1, x_407); -x_409 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_409 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_410 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_406); x_411 = lean_alloc_ctor(1, 3, 0); @@ -9573,7 +9316,7 @@ lean_inc(x_440); x_442 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_442, 0, x_440); lean_ctor_set(x_442, 1, x_441); -x_443 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_443 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_444 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_440); x_445 = lean_alloc_ctor(1, 3, 0); @@ -9728,7 +9471,7 @@ lean_inc(x_495); x_497 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_497, 0, x_495); lean_ctor_set(x_497, 1, x_496); -x_498 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_498 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_499 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_495); x_500 = lean_alloc_ctor(1, 3, 0); @@ -9792,7 +9535,7 @@ lean_inc(x_523); x_525 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_525, 0, x_523); lean_ctor_set(x_525, 1, x_524); -x_526 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_526 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_527 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_523); x_528 = lean_alloc_ctor(1, 3, 0); @@ -9869,7 +9612,7 @@ lean_inc(x_554); x_556 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_556, 0, x_554); lean_ctor_set(x_556, 1, x_555); -x_557 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_557 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_558 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_554); x_559 = lean_alloc_ctor(1, 3, 0); @@ -9963,7 +9706,7 @@ lean_inc(x_588); x_590 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_590, 0, x_588); lean_ctor_set(x_590, 1, x_589); -x_591 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_591 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_592 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_588); x_593 = lean_alloc_ctor(1, 3, 0); @@ -10147,7 +9890,7 @@ lean_inc(x_655); x_657 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_657, 0, x_655); lean_ctor_set(x_657, 1, x_656); -x_658 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_658 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_659 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_655); x_660 = lean_alloc_ctor(1, 3, 0); @@ -10211,7 +9954,7 @@ lean_inc(x_683); x_685 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_685, 0, x_683); lean_ctor_set(x_685, 1, x_684); -x_686 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_686 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_687 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_683); x_688 = lean_alloc_ctor(1, 3, 0); @@ -10288,7 +10031,7 @@ lean_inc(x_714); x_716 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_716, 0, x_714); lean_ctor_set(x_716, 1, x_715); -x_717 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_717 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_718 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_714); x_719 = lean_alloc_ctor(1, 3, 0); @@ -10382,7 +10125,7 @@ lean_inc(x_748); x_750 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_750, 0, x_748); lean_ctor_set(x_750, 1, x_749); -x_751 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_751 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_752 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_748); x_753 = lean_alloc_ctor(1, 3, 0); @@ -10566,7 +10309,7 @@ lean_inc(x_806); x_808 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_808, 0, x_806); lean_ctor_set(x_808, 1, x_807); -x_809 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_809 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_810 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_806); x_811 = lean_alloc_ctor(1, 3, 0); @@ -10630,7 +10373,7 @@ lean_inc(x_834); x_836 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_836, 0, x_834); lean_ctor_set(x_836, 1, x_835); -x_837 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_837 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_838 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_834); x_839 = lean_alloc_ctor(1, 3, 0); @@ -10707,7 +10450,7 @@ lean_inc(x_865); x_867 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_867, 0, x_865); lean_ctor_set(x_867, 1, x_866); -x_868 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_868 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_869 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_865); x_870 = lean_alloc_ctor(1, 3, 0); @@ -10801,7 +10544,7 @@ lean_inc(x_899); x_901 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_901, 0, x_899); lean_ctor_set(x_901, 1, x_900); -x_902 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_902 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_903 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_899); x_904 = lean_alloc_ctor(1, 3, 0); @@ -10986,7 +10729,7 @@ lean_inc(x_972); x_974 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_974, 0, x_972); lean_ctor_set(x_974, 1, x_973); -x_975 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_975 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_976 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_972); x_977 = lean_alloc_ctor(1, 3, 0); @@ -11050,7 +10793,7 @@ lean_inc(x_1000); x_1002 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1002, 0, x_1000); lean_ctor_set(x_1002, 1, x_1001); -x_1003 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1003 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1004 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1000); x_1005 = lean_alloc_ctor(1, 3, 0); @@ -11127,7 +10870,7 @@ lean_inc(x_1031); x_1033 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1033, 0, x_1031); lean_ctor_set(x_1033, 1, x_1032); -x_1034 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1034 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1035 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1031); x_1036 = lean_alloc_ctor(1, 3, 0); @@ -11221,7 +10964,7 @@ lean_inc(x_1065); x_1067 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1067, 0, x_1065); lean_ctor_set(x_1067, 1, x_1066); -x_1068 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1068 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1069 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1065); x_1070 = lean_alloc_ctor(1, 3, 0); @@ -11465,7 +11208,7 @@ lean_inc(x_1121); x_1123 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1123, 0, x_1121); lean_ctor_set(x_1123, 1, x_1122); -x_1124 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1124 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1125 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1121); x_1126 = lean_alloc_ctor(1, 3, 0); @@ -11529,7 +11272,7 @@ lean_inc(x_1149); x_1151 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1151, 0, x_1149); lean_ctor_set(x_1151, 1, x_1150); -x_1152 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1152 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1153 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1149); x_1154 = lean_alloc_ctor(1, 3, 0); @@ -11606,7 +11349,7 @@ lean_inc(x_1180); x_1182 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1182, 0, x_1180); lean_ctor_set(x_1182, 1, x_1181); -x_1183 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1183 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1184 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1180); x_1185 = lean_alloc_ctor(1, 3, 0); @@ -11700,7 +11443,7 @@ lean_inc(x_1214); x_1216 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1216, 0, x_1214); lean_ctor_set(x_1216, 1, x_1215); -x_1217 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1217 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1218 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1214); x_1219 = lean_alloc_ctor(1, 3, 0); @@ -11941,7 +11684,7 @@ lean_inc(x_1317); x_1319 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1319, 0, x_1317); lean_ctor_set(x_1319, 1, x_1318); -x_1320 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1320 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1321 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1317); x_1322 = lean_alloc_ctor(1, 3, 0); @@ -12005,7 +11748,7 @@ lean_inc(x_1345); x_1347 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1347, 0, x_1345); lean_ctor_set(x_1347, 1, x_1346); -x_1348 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1348 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1349 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1345); x_1350 = lean_alloc_ctor(1, 3, 0); @@ -12082,7 +11825,7 @@ lean_inc(x_1376); x_1378 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1378, 0, x_1376); lean_ctor_set(x_1378, 1, x_1377); -x_1379 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1379 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1380 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1376); x_1381 = lean_alloc_ctor(1, 3, 0); @@ -12176,7 +11919,7 @@ lean_inc(x_1410); x_1412 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1412, 0, x_1410); lean_ctor_set(x_1412, 1, x_1411); -x_1413 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1413 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1414 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1410); x_1415 = lean_alloc_ctor(1, 3, 0); @@ -12329,7 +12072,7 @@ lean_inc(x_1463); x_1465 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1465, 0, x_1463); lean_ctor_set(x_1465, 1, x_1464); -x_1466 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1466 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1467 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1463); x_1468 = lean_alloc_ctor(1, 3, 0); @@ -12393,7 +12136,7 @@ lean_inc(x_1491); x_1493 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1493, 0, x_1491); lean_ctor_set(x_1493, 1, x_1492); -x_1494 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1494 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1495 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1491); x_1496 = lean_alloc_ctor(1, 3, 0); @@ -12470,7 +12213,7 @@ lean_inc(x_1522); x_1524 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1524, 0, x_1522); lean_ctor_set(x_1524, 1, x_1523); -x_1525 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1525 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1526 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1522); x_1527 = lean_alloc_ctor(1, 3, 0); @@ -12564,7 +12307,7 @@ lean_inc(x_1556); x_1558 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1558, 0, x_1556); lean_ctor_set(x_1558, 1, x_1557); -x_1559 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1559 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1560 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1556); x_1561 = lean_alloc_ctor(1, 3, 0); @@ -12716,7 +12459,7 @@ lean_inc(x_1609); x_1611 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1611, 0, x_1609); lean_ctor_set(x_1611, 1, x_1610); -x_1612 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1612 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1613 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1609); x_1614 = lean_alloc_ctor(1, 3, 0); @@ -12780,7 +12523,7 @@ lean_inc(x_1637); x_1639 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1639, 0, x_1637); lean_ctor_set(x_1639, 1, x_1638); -x_1640 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1640 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1641 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1637); x_1642 = lean_alloc_ctor(1, 3, 0); @@ -12857,7 +12600,7 @@ lean_inc(x_1668); x_1670 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1670, 0, x_1668); lean_ctor_set(x_1670, 1, x_1669); -x_1671 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1671 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1672 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1668); x_1673 = lean_alloc_ctor(1, 3, 0); @@ -12951,7 +12694,7 @@ lean_inc(x_1702); x_1704 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1704, 0, x_1702); lean_ctor_set(x_1704, 1, x_1703); -x_1705 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1705 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1706 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1702); x_1707 = lean_alloc_ctor(1, 3, 0); @@ -13102,7 +12845,7 @@ lean_inc(x_1755); x_1757 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1757, 0, x_1755); lean_ctor_set(x_1757, 1, x_1756); -x_1758 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1758 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1759 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1755); x_1760 = lean_alloc_ctor(1, 3, 0); @@ -13166,7 +12909,7 @@ lean_inc(x_1783); x_1785 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1785, 0, x_1783); lean_ctor_set(x_1785, 1, x_1784); -x_1786 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1786 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1787 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1783); x_1788 = lean_alloc_ctor(1, 3, 0); @@ -13243,7 +12986,7 @@ lean_inc(x_1814); x_1816 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1816, 0, x_1814); lean_ctor_set(x_1816, 1, x_1815); -x_1817 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1817 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1818 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1814); x_1819 = lean_alloc_ctor(1, 3, 0); @@ -13337,7 +13080,7 @@ lean_inc(x_1848); x_1850 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1850, 0, x_1848); lean_ctor_set(x_1850, 1, x_1849); -x_1851 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1851 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1852 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1848); x_1853 = lean_alloc_ctor(1, 3, 0); @@ -13488,7 +13231,7 @@ lean_inc(x_1901); x_1903 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1903, 0, x_1901); lean_ctor_set(x_1903, 1, x_1902); -x_1904 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1904 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1905 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1901); x_1906 = lean_alloc_ctor(1, 3, 0); @@ -13552,7 +13295,7 @@ lean_inc(x_1929); x_1931 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1931, 0, x_1929); lean_ctor_set(x_1931, 1, x_1930); -x_1932 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1932 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1933 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1929); x_1934 = lean_alloc_ctor(1, 3, 0); @@ -13629,7 +13372,7 @@ lean_inc(x_1960); x_1962 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1962, 0, x_1960); lean_ctor_set(x_1962, 1, x_1961); -x_1963 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1963 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1964 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1960); x_1965 = lean_alloc_ctor(1, 3, 0); @@ -13723,7 +13466,7 @@ lean_inc(x_1994); x_1996 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_1996, 0, x_1994); lean_ctor_set(x_1996, 1, x_1995); -x_1997 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_1997 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_1998 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_1994); x_1999 = lean_alloc_ctor(1, 3, 0); @@ -15888,12 +15631,12 @@ lean_ctor_set(x_27, 1, x_26); x_28 = l_Lean_Elab_Term_expandWhereDecls___closed__8; lean_inc(x_23); x_29 = l_Lean_Syntax_node2(x_23, x_28, x_25, x_27); -x_30 = l_Lean_Elab_Term_expandForall___closed__5; +x_30 = l_Lean_Elab_Term_expandForall___closed__8; x_31 = l_Lean_Syntax_SepArray_ofElems(x_30, x_20); lean_dec(x_20); x_32 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_33 = l_Array_append___rarg(x_32, x_31); -x_34 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_34 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_23); x_35 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_35, 0, x_23); @@ -16159,7 +15902,7 @@ static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expand { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(2); -x_2 = l_Lean_Elab_Term_expandForall___closed__5; +x_2 = l_Lean_Elab_Term_expandForall___closed__8; x_3 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -16251,7 +15994,7 @@ lean_inc(x_21); x_37 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_37, 0, x_21); lean_ctor_set(x_37, 1, x_36); -x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_39 = l_Lean_Syntax_node1(x_21, x_38, x_26); x_40 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16286,7 +16029,7 @@ lean_inc(x_21); x_51 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_51, 0, x_21); lean_ctor_set(x_51, 1, x_50); -x_52 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_52 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_53 = l_Lean_Syntax_node1(x_21, x_52, x_26); x_54 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16324,7 +16067,7 @@ lean_inc(x_21); x_66 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_66, 0, x_21); lean_ctor_set(x_66, 1, x_65); -x_67 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_67 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_68 = l_Lean_Syntax_node1(x_21, x_67, x_26); x_69 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16361,7 +16104,7 @@ 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_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_82 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_83 = l_Lean_Syntax_node1(x_21, x_82, x_26); x_84 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16403,7 +16146,7 @@ lean_inc(x_21); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_21); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_98 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_99 = l_Lean_Syntax_node1(x_21, x_98, x_26); x_100 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -16434,7 +16177,7 @@ lean_inc(x_21); x_110 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_110, 0, x_21); lean_ctor_set(x_110, 1, x_109); -x_111 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_111 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_21); x_112 = l_Lean_Syntax_node1(x_21, x_111, x_26); x_113 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -16530,7 +16273,7 @@ lean_inc(x_128); x_145 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_145, 0, x_128); lean_ctor_set(x_145, 1, x_144); -x_146 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_146 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_128); x_147 = l_Lean_Syntax_node1(x_128, x_146, x_133); x_148 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16578,7 +16321,7 @@ lean_inc(x_128); x_161 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_161, 0, x_128); lean_ctor_set(x_161, 1, x_160); -x_162 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_162 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_128); x_163 = l_Lean_Syntax_node1(x_128, x_162, x_133); x_164 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16630,7 +16373,7 @@ lean_inc(x_128); x_178 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_178, 0, x_128); lean_ctor_set(x_178, 1, x_177); -x_179 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_179 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_128); x_180 = l_Lean_Syntax_node1(x_128, x_179, x_133); x_181 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -16754,7 +16497,7 @@ lean_inc(x_201); x_218 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_218, 0, x_201); lean_ctor_set(x_218, 1, x_217); -x_219 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_219 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_201); x_220 = l_Lean_Syntax_node1(x_201, x_219, x_206); x_221 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16802,7 +16545,7 @@ lean_inc(x_201); x_234 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_234, 0, x_201); lean_ctor_set(x_234, 1, x_233); -x_235 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_235 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_201); x_236 = l_Lean_Syntax_node1(x_201, x_235, x_206); x_237 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -16854,7 +16597,7 @@ lean_inc(x_201); x_251 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_251, 0, x_201); lean_ctor_set(x_251, 1, x_250); -x_252 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_252 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_201); x_253 = l_Lean_Syntax_node1(x_201, x_252, x_206); x_254 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__13; @@ -16894,7 +16637,7 @@ lean_inc(x_264); x_266 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_266, 0, x_264); lean_ctor_set(x_266, 1, x_265); -x_267 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_267 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_268 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_264); x_269 = lean_alloc_ctor(1, 3, 0); @@ -16943,7 +16686,7 @@ lean_inc(x_285); x_287 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_287, 0, x_285); lean_ctor_set(x_287, 1, x_286); -x_288 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_288 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_289 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_285); x_290 = lean_alloc_ctor(1, 3, 0); @@ -17247,7 +16990,7 @@ lean_inc(x_19); x_32 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_32, 0, x_19); lean_ctor_set(x_32, 1, x_31); -x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_33 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_19); x_34 = l_Lean_Syntax_node1(x_19, x_33, x_24); x_35 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -17290,7 +17033,7 @@ lean_inc(x_19); x_50 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_50, 0, x_19); lean_ctor_set(x_50, 1, x_49); -x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_19); x_52 = l_Lean_Syntax_node1(x_19, x_51, x_24); x_53 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -17408,7 +17151,7 @@ lean_inc(x_75); x_89 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_89, 0, x_75); lean_ctor_set(x_89, 1, x_88); -x_90 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_90 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_75); x_91 = l_Lean_Syntax_node1(x_75, x_90, x_80); x_92 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -17556,7 +17299,7 @@ lean_inc(x_119); x_133 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_133, 0, x_119); lean_ctor_set(x_133, 1, x_132); -x_134 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_134 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_119); x_135 = l_Lean_Syntax_node1(x_119, x_134, x_124); x_136 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; @@ -17627,7 +17370,7 @@ lean_inc(x_153); x_155 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_155, 0, x_153); lean_ctor_set(x_155, 1, x_154); -x_156 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_156 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_157 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_153); x_158 = lean_alloc_ctor(1, 3, 0); @@ -17982,7 +17725,7 @@ lean_ctor_set(x_47, 0, x_45); lean_ctor_set(x_47, 1, x_46); x_48 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_49 = l_Array_append___rarg(x_48, x_41); -x_50 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_50 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_45); x_51 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_51, 0, x_45); @@ -18028,7 +17771,7 @@ lean_ctor_set(x_64, 0, x_62); lean_ctor_set(x_64, 1, x_63); x_65 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_66 = l_Array_append___rarg(x_65, x_58); -x_67 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_67 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_62); x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_62); @@ -18114,7 +17857,7 @@ x_88 = lean_array_get_size(x_87); x_89 = lean_usize_of_nat(x_88); lean_dec(x_88); x_90 = 0; -x_91 = l_Lean_Elab_Term_quoteAutoTactic___closed__11; +x_91 = l_Lean_Elab_Term_expandForall___closed__7; lean_inc(x_2); x_92 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFun___spec__1(x_91, x_84, x_89, x_90, x_87, x_2, x_3); if (lean_obj_tag(x_92) == 0) @@ -18137,7 +17880,7 @@ lean_ctor_set(x_99, 0, x_97); lean_ctor_set(x_99, 1, x_98); x_100 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_101 = l_Array_append___rarg(x_100, x_94); -x_102 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_102 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_97); x_103 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_103, 0, x_97); @@ -18179,7 +17922,7 @@ lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); x_116 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; x_117 = l_Array_append___rarg(x_116, x_109); -x_118 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_118 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; lean_inc(x_113); x_119 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_119, 0, x_113); @@ -18260,7 +18003,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_expandFun___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -18291,7 +18034,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(591u); +x_1 = lean_unsigned_to_nat(595u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18303,7 +18046,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(602u); +x_1 = lean_unsigned_to_nat(606u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18331,7 +18074,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(591u); +x_1 = lean_unsigned_to_nat(595u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18343,7 +18086,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(591u); +x_1 = lean_unsigned_to_nat(595u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18471,7 +18214,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun___clos { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_expandExplicitFun___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -18502,7 +18245,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(604u); +x_1 = lean_unsigned_to_nat(608u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18514,7 +18257,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _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(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18542,7 +18285,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(604u); +x_1 = lean_unsigned_to_nat(608u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18554,7 +18297,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(604u); +x_1 = lean_unsigned_to_nat(608u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19827,7 +19570,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_precheckFun___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_precheckFun___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -20508,7 +20251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabFun___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -20539,7 +20282,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(621u); +x_1 = lean_unsigned_to_nat(625u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20551,7 +20294,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(634u); +x_1 = lean_unsigned_to_nat(638u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20579,7 +20322,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(621u); +x_1 = lean_unsigned_to_nat(625u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20591,7 +20334,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(621u); +x_1 = lean_unsigned_to_nat(625u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22053,7 +21796,7 @@ lean_inc(x_21); x_22 = lean_ctor_get(x_20, 1); lean_inc(x_22); lean_dec(x_20); -x_23 = l_Lean_Elab_Term_quoteAutoTactic___closed__44; +x_23 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9; x_24 = lean_array_push(x_23, x_5); x_25 = 0; x_26 = 1; @@ -22260,32 +22003,24 @@ return x_40; static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("decl", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_2 = l_Lean_Elab_Term_expandWhereDecls___closed__9; -x_3 = l_Lean_Elab_Term_elabLetDeclAux___closed__1; +x_3 = l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__3() { _start: { lean_object* x_1; @@ -22293,16 +22028,16 @@ x_1 = lean_mk_string_from_bytes(" : ", 3); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_elabLetDeclAux___closed__4; +x_1 = l_Lean_Elab_Term_elabLetDeclAux___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__5() { _start: { lean_object* x_1; @@ -22310,11 +22045,11 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_elabLetDeclAux___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_elabLetDeclAux___closed__6; +x_1 = l_Lean_Elab_Term_elabLetDeclAux___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -22357,7 +22092,7 @@ lean_dec(x_21); x_26 = l_Lean_Syntax_getId(x_1); lean_inc(x_26); x_27 = l_Lean_Elab_Term_kindOfBinderName(x_26); -x_28 = l_Lean_Elab_Term_elabLetDeclAux___closed__2; +x_28 = l_Lean_Elab_Term_elabLetDeclAux___closed__1; x_29 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_28, x_10, x_11, x_12, x_13, x_14, x_15, x_22); x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); @@ -22381,11 +22116,11 @@ lean_inc(x_35); lean_dec(x_29); lean_inc(x_26); x_36 = l_Lean_MessageData_ofName(x_26); -x_37 = l_Lean_Elab_Term_elabLetDeclAux___closed__3; +x_37 = l_Lean_Elab_Term_elabLetDeclAux___closed__2; x_38 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_38, 0, x_37); lean_ctor_set(x_38, 1, x_36); -x_39 = l_Lean_Elab_Term_elabLetDeclAux___closed__5; +x_39 = l_Lean_Elab_Term_elabLetDeclAux___closed__4; x_40 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); @@ -22394,7 +22129,7 @@ x_41 = l_Lean_MessageData_ofExpr(x_23); x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_Elab_Term_elabLetDeclAux___closed__7; +x_43 = l_Lean_Elab_Term_elabLetDeclAux___closed__6; x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); @@ -22660,7 +22395,7 @@ x_12 = lean_unsigned_to_nat(1u); x_13 = l_Lean_Syntax_getArg(x_1, x_12); x_14 = lean_unsigned_to_nat(2u); x_15 = l_Lean_Syntax_getArg(x_1, x_14); -x_16 = l_Lean_Elab_Term_elabLetDeclAux___closed__6; +x_16 = l_Lean_Elab_Term_elabLetDeclAux___closed__5; x_17 = 0; x_18 = l_Lean_mkAtomFrom(x_1, x_16, x_17); x_19 = l_Lean_Elab_Term_expandLetEqnsDecl___closed__3; @@ -22692,7 +22427,7 @@ x_32 = lean_unsigned_to_nat(1u); x_33 = l_Lean_Syntax_getArg(x_1, x_32); x_34 = lean_unsigned_to_nat(2u); x_35 = l_Lean_Syntax_getArg(x_1, x_34); -x_36 = l_Lean_Elab_Term_elabLetDeclAux___closed__6; +x_36 = l_Lean_Elab_Term_elabLetDeclAux___closed__5; x_37 = 0; x_38 = l_Lean_mkAtomFrom(x_1, x_36, x_37); x_39 = l_Lean_Elab_Term_expandLetEqnsDecl___closed__3; @@ -22879,7 +22614,7 @@ lean_inc(x_31); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_31); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_36 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_37 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_31); x_38 = lean_alloc_ctor(1, 3, 0); @@ -22962,7 +22697,7 @@ lean_inc(x_69); x_73 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_73, 0, x_69); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9; +x_74 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; x_75 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux___rarg___closed__1; lean_inc(x_69); x_76 = lean_alloc_ctor(1, 3, 0); @@ -23472,7 +23207,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__3( { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__2; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -23503,7 +23238,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(746u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23515,7 +23250,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(747u); +x_1 = lean_unsigned_to_nat(751u); x_2 = lean_unsigned_to_nat(129u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23543,7 +23278,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(746u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23555,7 +23290,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(746u); +x_1 = lean_unsigned_to_nat(750u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23643,7 +23378,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed_ { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetFunDecl___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -23674,7 +23409,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(749u); +x_1 = lean_unsigned_to_nat(753u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23686,7 +23421,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(750u); +x_1 = lean_unsigned_to_nat(754u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23714,7 +23449,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(749u); +x_1 = lean_unsigned_to_nat(753u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23726,7 +23461,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(749u); +x_1 = lean_unsigned_to_nat(753u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23815,7 +23550,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___clo { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -23846,7 +23581,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(752u); +x_1 = lean_unsigned_to_nat(756u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23858,7 +23593,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(757u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23886,7 +23621,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(752u); +x_1 = lean_unsigned_to_nat(756u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23898,7 +23633,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(752u); +x_1 = lean_unsigned_to_nat(756u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23987,7 +23722,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed_ { 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_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_4 = l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); @@ -24018,7 +23753,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(755u); +x_1 = lean_unsigned_to_nat(759u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24030,7 +23765,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(756u); +x_1 = lean_unsigned_to_nat(760u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24058,7 +23793,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(755u); +x_1 = lean_unsigned_to_nat(759u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24070,7 +23805,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(755u); +x_1 = lean_unsigned_to_nat(759u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24116,17 +23851,17 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_1 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_2 = l_Lean_Elab_Term_expandWhereDecls___closed__9; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -24136,27 +23871,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5() { _start: { lean_object* x_1; @@ -24164,17 +23899,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7() { _start: { lean_object* x_1; @@ -24182,37 +23917,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9; -x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9; +x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11() { _start: { lean_object* x_1; @@ -24220,17 +23955,17 @@ x_1 = lean_mk_string_from_bytes("Binders", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13() { _start: { lean_object* x_1; @@ -24238,33 +23973,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14; -x_2 = lean_unsigned_to_nat(10339u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14; +x_2 = lean_unsigned_to_nat(10159u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -24310,28 +24045,62 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__2 = _in lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__2); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__3 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__3(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__6); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__7); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__8); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__9); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__10); -l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__1___closed__11); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__7); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__1); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__2); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__3); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__4); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__5); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__6); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__7 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__7(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__7); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__9); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__10); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__13); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__18); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__19); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20); +l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__21 = _init_l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__21(); +lean_mark_persistent(l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__21); l_Lean_Elab_Term_quoteAutoTactic___closed__1 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__1); l_Lean_Elab_Term_quoteAutoTactic___closed__2 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__2(); @@ -24390,68 +24159,12 @@ l_Lean_Elab_Term_quoteAutoTactic___closed__28 = _init_l_Lean_Elab_Term_quoteAuto lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__28); l_Lean_Elab_Term_quoteAutoTactic___closed__29 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__29(); lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__29); -l_Lean_Elab_Term_quoteAutoTactic___closed__30 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__30(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__30); -l_Lean_Elab_Term_quoteAutoTactic___closed__31 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__31(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__31); -l_Lean_Elab_Term_quoteAutoTactic___closed__32 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__32(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__32); -l_Lean_Elab_Term_quoteAutoTactic___closed__33 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__33(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__33); -l_Lean_Elab_Term_quoteAutoTactic___closed__34 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__34(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__34); -l_Lean_Elab_Term_quoteAutoTactic___closed__35 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__35(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__35); -l_Lean_Elab_Term_quoteAutoTactic___closed__36 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__36(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__36); -l_Lean_Elab_Term_quoteAutoTactic___closed__37 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__37(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__37); -l_Lean_Elab_Term_quoteAutoTactic___closed__38 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__38(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__38); -l_Lean_Elab_Term_quoteAutoTactic___closed__39 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__39(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__39); -l_Lean_Elab_Term_quoteAutoTactic___closed__40 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__40(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__40); -l_Lean_Elab_Term_quoteAutoTactic___closed__41 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__41(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__41); -l_Lean_Elab_Term_quoteAutoTactic___closed__42 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__42(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__42); -l_Lean_Elab_Term_quoteAutoTactic___closed__43 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__43(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__43); -l_Lean_Elab_Term_quoteAutoTactic___closed__44 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__44(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__44); -l_Lean_Elab_Term_quoteAutoTactic___closed__45 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__45(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__45); -l_Lean_Elab_Term_quoteAutoTactic___closed__46 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__46(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__46); -l_Lean_Elab_Term_quoteAutoTactic___closed__47 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__47(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__47); -l_Lean_Elab_Term_quoteAutoTactic___closed__48 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__48(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__48); -l_Lean_Elab_Term_quoteAutoTactic___closed__49 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__49(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__49); -l_Lean_Elab_Term_quoteAutoTactic___closed__50 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__50(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__50); -l_Lean_Elab_Term_quoteAutoTactic___closed__51 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__51(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__51); -l_Lean_Elab_Term_quoteAutoTactic___closed__52 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__52(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__52); -l_Lean_Elab_Term_quoteAutoTactic___closed__53 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__53(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__53); -l_Lean_Elab_Term_quoteAutoTactic___closed__54 = _init_l_Lean_Elab_Term_quoteAutoTactic___closed__54(); -lean_mark_persistent(l_Lean_Elab_Term_quoteAutoTactic___closed__54); l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1); l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__2); l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3(); lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__3); -l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__4); -l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__5); -l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6 = _init_l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__6); l_Lean_Elab_Term_declareTacticSyntax___closed__1 = _init_l_Lean_Elab_Term_declareTacticSyntax___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_declareTacticSyntax___closed__1); l_Lean_Elab_Term_declareTacticSyntax___closed__2 = _init_l_Lean_Elab_Term_declareTacticSyntax___closed__2(); @@ -24482,14 +24195,16 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__1 lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__12); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__13); +l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__14); +l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15); l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__1); l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__2); l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3(); lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__3); -l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4 = _init_l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_getBinderIds___spec__1___closed__4); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__1 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__1); l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__2 = _init_l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__2(); @@ -24510,6 +24225,8 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__7 = _ini lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__7); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__8 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__8(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__8); +l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__9); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__1 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__1); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__2 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__2(); @@ -24524,19 +24241,19 @@ l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed_ lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__3); l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4 = _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4(); lean_mark_persistent(l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967____closed__6); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1967_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__6); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Term_checkBinderAnnotations = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Term_checkBinderAnnotations); @@ -24579,6 +24296,12 @@ l_Lean_Elab_Term_expandForall___closed__4 = _init_l_Lean_Elab_Term_expandForall_ lean_mark_persistent(l_Lean_Elab_Term_expandForall___closed__4); l_Lean_Elab_Term_expandForall___closed__5 = _init_l_Lean_Elab_Term_expandForall___closed__5(); lean_mark_persistent(l_Lean_Elab_Term_expandForall___closed__5); +l_Lean_Elab_Term_expandForall___closed__6 = _init_l_Lean_Elab_Term_expandForall___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_expandForall___closed__6); +l_Lean_Elab_Term_expandForall___closed__7 = _init_l_Lean_Elab_Term_expandForall___closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_expandForall___closed__7); +l_Lean_Elab_Term_expandForall___closed__8 = _init_l_Lean_Elab_Term_expandForall___closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_expandForall___closed__8); l___regBuiltin_Lean_Elab_Term_expandForall___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_expandForall___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_expandForall___closed__1); l___regBuiltin_Lean_Elab_Term_expandForall___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_expandForall___closed__2(); @@ -24918,8 +24641,6 @@ l_Lean_Elab_Term_elabLetDeclAux___closed__5 = _init_l_Lean_Elab_Term_elabLetDecl lean_mark_persistent(l_Lean_Elab_Term_elabLetDeclAux___closed__5); l_Lean_Elab_Term_elabLetDeclAux___closed__6 = _init_l_Lean_Elab_Term_elabLetDeclAux___closed__6(); lean_mark_persistent(l_Lean_Elab_Term_elabLetDeclAux___closed__6); -l_Lean_Elab_Term_elabLetDeclAux___closed__7 = _init_l_Lean_Elab_Term_elabLetDeclAux___closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_elabLetDeclAux___closed__7); l_Lean_Elab_Term_expandLetEqnsDecl___closed__1 = _init_l_Lean_Elab_Term_expandLetEqnsDecl___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_expandLetEqnsDecl___closed__1); l_Lean_Elab_Term_expandLetEqnsDecl___closed__2 = _init_l_Lean_Elab_Term_expandLetEqnsDecl___closed__2(); @@ -25060,37 +24781,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___cl if (builtin) {res = l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10339_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(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/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index 797369074d95..47afc0c4e0e8 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -74,7 +74,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_d lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__19___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Term_elabUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandSuffices_declRange___closed__2; @@ -88,7 +88,6 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f_expandCDotArg_x3f(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__3; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__14; static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__7; @@ -103,12 +102,13 @@ static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymousCto static lean_object* l_Lean_Elab_Term_elabShow___closed__4; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_withLocalIdentFor___closed__1; static lean_object* l_Lean_Elab_Term_expandHave___lambda__2___closed__3; +static lean_object* l_Lean_Elab_Term_elabSubst___lambda__6___closed__2; lean_object* l_Lean_Meta_mkEqSymm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__2; @@ -264,7 +264,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry___closed__5; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__11; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabAnonymousCtor_declRange___closed__3; -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__4; static lean_object* l_Lean_Elab_Term_elabSubst___closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__15; @@ -283,7 +282,6 @@ static lean_object* l_Lean_Elab_Term_expandHave___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__17; LEAN_EXPORT lean_object* l_Lean_exprDependsOn___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTypeAscription_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex___closed__4; static lean_object* l_Lean_Elab_Term_elabSorry___closed__7; @@ -466,7 +464,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__ static lean_object* l___regBuiltin_Lean_Elab_Term_expandHave___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoeFunNotation_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabCDotFunctionAlias_x3f___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_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___regBuiltin_Lean_Elab_Term_expandUnreachable_declRange___closed__7; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__36; @@ -522,12 +520,13 @@ static lean_object* l_Lean_Elab_Term_expandTuple___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_Term_elabSubst___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic_declRange___closed__7; extern lean_object* l_Lean_Elab_Term_termElabAttribute; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabTrailingParserMacro___lambda__3___closed__2; uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabHaveI_declRange(lean_object*); @@ -577,7 +576,7 @@ lean_object* l_Lean_throwError___at___private_Lean_Elab_Term_0__Lean_Elab_Term_e static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__7; static lean_object* l_Lean_Elab_Term_elabShow___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_Term_elabSubst___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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription_declRange___closed__3; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__21; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCoe___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*); @@ -665,7 +664,7 @@ static lean_object* l_Lean_Elab_Term_elabPanic___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__4(lean_object*, size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabPanic___closed__3; static lean_object* l_Lean_Elab_Term_expandUnreachable___rarg___closed__4; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__7; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -693,7 +692,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription___closed__1 static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabRunElab___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__9; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_Elab_Term_elabSubst___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* l_Array_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__2(lean_object*); @@ -772,7 +771,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term LEAN_EXPORT lean_object* l_Lean_setImplementedBy___at_Lean_Elab_Term_elabUnsafe___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandSuffices___closed__9; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__16(lean_object*, lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandHave___lambda__4___closed__1; static lean_object* l_Lean_Elab_Term_expandTuple___closed__7; static lean_object* l_Lean_Elab_Term_elabAnonymousCtor___closed__4; @@ -836,7 +835,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandSuffices___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__2; static lean_object* l_Lean_Elab_Term_expandHave___closed__15; static lean_object* l_Lean_Elab_Term_expandUnreachable___rarg___closed__2; static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__33; @@ -849,7 +847,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___close static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandCDot_x3f_go___spec__3___closed__2; -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__6; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed___closed__2; @@ -862,6 +859,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___clos LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandUnreachable___closed__3; +static lean_object* l_Lean_Elab_Term_elabSubst___lambda__6___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___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_Elab_BuiltinNotation_0__Lean_Elab_Term_elabTParserMacroAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object*); @@ -885,6 +883,7 @@ lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabRunElab___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabShow___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__4; +static lean_object* l_Lean_Elab_Term_elabSubst___lambda__6___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabHaveI___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabStateRefT___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -917,6 +916,7 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandHave___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandParen___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_hasCDot___spec__1___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_elabSubst___lambda__6___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__7; lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabSubst___closed__13; @@ -963,6 +963,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabAnonymousCtor___lambda__1___boxed( lean_object* l_String_intercalate(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandCDot_x3f_go___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabAnonymousCtor___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_elabSubst___lambda__6___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst___closed__3; uint8_t l_Lean_Syntax_structEq(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex___closed__1; @@ -977,7 +978,7 @@ static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elab lean_object* l_Lean_Elab_Term_evalTerm___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_getMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandAssert___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__5; uint8_t l_Lean_Expr_hasFVar(lean_object*); @@ -991,7 +992,6 @@ static lean_object* l_Lean_Elab_Term_elabRunElab___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandTypeAscription_declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_isSubstCandidate___spec__12(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabRunElab___closed__6; -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__6; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__24; @@ -1038,6 +1038,7 @@ static lean_object* l_Lean_Elab_Term_elabPanic___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange___closed__2; static lean_object* l_Lean_Elab_Term_expandAssert___closed__6; +static lean_object* l_Lean_Elab_Term_elabSubst___lambda__6___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe_declRange___closed__6; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandAssert_declRange___closed__3; @@ -1069,7 +1070,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange(l static lean_object* l___regBuiltin_Lean_Elab_Term_elabLeadingParserMacro___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLeadingParserMacro(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoe___closed__4; -static lean_object* l_Lean_Elab_Term_elabSubst___lambda__5___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabBorrowed_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___closed__1; lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__1(lean_object*, lean_object*); @@ -1107,7 +1107,7 @@ LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x static lean_object* l_Lean_Elab_Term_elabRunElab_unsafe__1___closed__9; lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLeadingParserMacro___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__10___boxed(lean_object**); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__9; static lean_object* l_Lean_Elab_Term_expandHave___closed__13; lean_object* l_Lean_Meta_mkOfNonempty(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -22058,132 +22058,185 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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: { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; -x_14 = l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandCDot_x3f_go___spec__3___closed__2; -x_15 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_14, x_11, x_12, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; +x_15 = l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandCDot_x3f_go___spec__3___closed__2; +x_16 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_15, x_12, x_13, x_14); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__2), 10, 2); -lean_closure_set(x_18, 0, x_1); -lean_closure_set(x_18, 1, x_2); -x_19 = 0; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__2), 10, 2); +lean_closure_set(x_19, 0, x_5); +lean_closure_set(x_19, 1, x_3); x_20 = 0; +x_21 = 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); -lean_inc(x_7); -x_21 = l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(x_16, x_19, x_3, x_18, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_17); -if (lean_obj_tag(x_21) == 0) +x_22 = l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(x_17, x_20, x_1, x_19, x_21, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +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_22); -x_24 = l_Lean_Meta_isTypeCorrect(x_22, x_9, x_10, x_11, x_12, x_23); -if (lean_obj_tag(x_24) == 0) +lean_inc(x_23); +x_25 = l_Lean_Meta_isTypeCorrect(x_23, x_10, x_11, x_12, x_13, x_24); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_25; uint8_t x_26; -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_unbox(x_25); -lean_dec(x_25); -if (x_26 == 0) +lean_object* x_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_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; -lean_dec(x_22); -lean_dec(x_5); +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +lean_dec(x_23); lean_dec(x_4); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__2; -x_29 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_28, x_7, x_8, x_9, x_10, x_11, x_12, x_27); +lean_dec(x_2); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = l_Lean_Elab_Term_elabSubst___lambda__4___closed__2; +x_30 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_29, x_8, x_9, x_10, x_11, x_12, x_13, x_28); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -x_30 = !lean_is_exclusive(x_29); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) { -return x_29; +return x_30; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_29, 0); -x_32 = lean_ctor_get(x_29, 1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_30, 0); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_29); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_dec(x_30); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } else { -lean_object* x_34; lean_object* x_35; +lean_object* x_35; lean_object* x_36; +lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -x_34 = lean_ctor_get(x_24, 1); -lean_inc(x_34); -lean_dec(x_24); -x_35 = l_Lean_Meta_mkEqRec(x_22, x_4, x_5, x_9, x_10, x_11, x_12, x_34); -return x_35; +x_35 = lean_ctor_get(x_25, 1); +lean_inc(x_35); +lean_dec(x_25); +x_36 = l_Lean_Meta_mkEqRec(x_23, x_2, x_4, x_10, x_11, x_12, x_13, x_35); +return x_36; } } else { -uint8_t x_36; -lean_dec(x_22); +uint8_t x_37; +lean_dec(x_23); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); lean_dec(x_4); -x_36 = !lean_is_exclusive(x_24); -if (x_36 == 0) +lean_dec(x_2); +x_37 = !lean_is_exclusive(x_25); +if (x_37 == 0) { -return x_24; +return x_25; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_24, 0); -x_38 = lean_ctor_get(x_24, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_25, 0); +x_39 = lean_ctor_get(x_25, 1); +lean_inc(x_39); lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_24); -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_dec(x_25); +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 +{ +uint8_t x_41; +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_4); +lean_dec(x_2); +x_41 = !lean_is_exclusive(x_22); +if (x_41 == 0) +{ +return x_22; +} +else +{ +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; +} } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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) { +_start: +{ +lean_object* x_14; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_14 = l_Lean_Meta_mkEqSymm(x_3, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +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(0); +x_18 = lean_apply_12(x_1, x_2, x_15, x_5, x_4, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_16); +return x_18; +} else { -uint8_t x_40; +uint8_t x_19; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -22192,28 +22245,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -x_40 = !lean_is_exclusive(x_21); -if (x_40 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_19 = !lean_is_exclusive(x_14); +if (x_19 == 0) { -return x_21; +return x_14; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_21, 0); -x_42 = lean_ctor_get(x_21, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_21); -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_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; } } } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -22221,19 +22276,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__2() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; -x_4 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__1; +x_4 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__3() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__3() { _start: { lean_object* x_1; @@ -22241,19 +22296,19 @@ x_1 = lean_mk_string_from_bytes("subst", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__4() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; -x_4 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__3; +x_4 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__5() { _start: { lean_object* x_1; @@ -22261,19 +22316,19 @@ x_1 = lean_mk_string_from_bytes("exact", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_3 = l_Lean_Elab_Term_expandSuffices___closed__7; -x_4 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__5; +x_4 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__5; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_Term_elabSubst___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) { _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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; 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; @@ -22290,7 +22345,7 @@ lean_inc(x_14); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_14); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__3; +x_19 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__3; lean_inc(x_14); x_20 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_20, 0, x_14); @@ -22298,7 +22353,7 @@ lean_ctor_set(x_20, 1, x_19); x_21 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; lean_inc(x_14); x_22 = l_Lean_Syntax_node1(x_14, x_21, x_1); -x_23 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__4; +x_23 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__4; lean_inc(x_14); x_24 = l_Lean_Syntax_node2(x_14, x_23, x_20, x_22); x_25 = l_Lean_Elab_Term_elabShow___closed__13; @@ -22306,17 +22361,17 @@ lean_inc(x_14); x_26 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_26, 0, x_14); lean_ctor_set(x_26, 1, x_25); -x_27 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__5; +x_27 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__5; lean_inc(x_14); x_28 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_28, 0, x_14); lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__6; +x_29 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__6; lean_inc(x_14); x_30 = l_Lean_Syntax_node2(x_14, x_29, x_28, x_4); lean_inc(x_14); x_31 = l_Lean_Syntax_node3(x_14, x_21, x_24, x_26, x_30); -x_32 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__2; +x_32 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__2; lean_inc(x_14); x_33 = l_Lean_Syntax_node1(x_14, x_32, x_31); x_34 = l_Lean_Elab_Term_expandSuffices___closed__9; @@ -22339,7 +22394,7 @@ x_43 = l_Lean_Elab_Term_withMacroExpansion___rarg(x_3, x_37, x_42, x_5, x_6, x_7 return x_43; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; @@ -22353,7 +22408,7 @@ x_16 = l_Lean_Expr_hasMVar(x_14); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__5), 11, 3); +x_17 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__6), 11, 3); lean_closure_set(x_17, 0, x_5); lean_closure_set(x_17, 1, x_2); lean_closure_set(x_17, 2, x_3); @@ -22377,7 +22432,7 @@ lean_inc(x_21); x_25 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_25, 0, x_21); lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__3; +x_26 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__3; lean_inc(x_21); x_27 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_27, 0, x_21); @@ -22385,7 +22440,7 @@ lean_ctor_set(x_27, 1, x_26); x_28 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__3; lean_inc(x_21); x_29 = l_Lean_Syntax_node1(x_21, x_28, x_5); -x_30 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__4; +x_30 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__4; lean_inc(x_21); x_31 = l_Lean_Syntax_node2(x_21, x_30, x_27, x_29); x_32 = l_Lean_Elab_Term_elabShow___closed__13; @@ -22393,17 +22448,17 @@ lean_inc(x_21); x_33 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_33, 0, x_21); lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__5; +x_34 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__5; lean_inc(x_21); x_35 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_35, 0, x_21); lean_ctor_set(x_35, 1, x_34); -x_36 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__6; +x_36 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__6; lean_inc(x_21); x_37 = l_Lean_Syntax_node2(x_21, x_36, x_35, x_4); lean_inc(x_21); x_38 = l_Lean_Syntax_node3(x_21, x_28, x_31, x_33, x_37); -x_39 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__2; +x_39 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__2; lean_inc(x_21); x_40 = l_Lean_Syntax_node1(x_21, x_39, x_38); x_41 = l_Lean_Elab_Term_expandSuffices___closed__9; @@ -22427,7 +22482,7 @@ return x_50; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object* x_14) { _start: { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; @@ -22677,7 +22732,7 @@ return x_63; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { lean_object* x_18; lean_object* x_19; @@ -22741,7 +22796,7 @@ x_26 = lean_ctor_get(x_19, 1); lean_inc(x_26); lean_dec(x_19); x_27 = lean_box(0); -x_28 = l_Lean_Elab_Term_elabSubst___lambda__7(x_5, x_1, x_6, x_7, x_8, x_9, x_27, x_11, x_12, x_13, x_14, x_15, x_16, x_26); +x_28 = l_Lean_Elab_Term_elabSubst___lambda__8(x_5, x_1, x_6, x_7, x_8, x_9, x_27, x_11, x_12, x_13, x_14, x_15, x_16, x_26); return x_28; } } @@ -22782,7 +22837,7 @@ return x_32; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18) { _start: { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_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; uint8_t x_81; uint8_t x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; @@ -22960,7 +23015,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_107 = l_Lean_Elab_Term_elabSubst___lambda__8(x_103, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_106, x_12, x_13, x_14, x_15, x_84, x_17, x_104); +x_107 = l_Lean_Elab_Term_elabSubst___lambda__9(x_103, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_106, x_12, x_13, x_14, x_15, x_84, x_17, x_104); if (lean_obj_tag(x_107) == 0) { lean_object* x_108; lean_object* x_109; @@ -23062,7 +23117,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_119 = l_Lean_Elab_Term_elabSubst___lambda__8(x_114, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_118, x_12, x_13, x_14, x_15, x_84, x_17, x_115); +x_119 = l_Lean_Elab_Term_elabSubst___lambda__9(x_114, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_118, x_12, x_13, x_14, x_15, x_84, x_17, x_115); if (lean_obj_tag(x_119) == 0) { lean_object* x_120; lean_object* x_121; @@ -23303,7 +23358,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_144 = l_Lean_Elab_Term_elabSubst___lambda__8(x_140, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_143, x_12, x_13, x_14, x_15, x_84, x_17, x_141); +x_144 = l_Lean_Elab_Term_elabSubst___lambda__9(x_140, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_143, x_12, x_13, x_14, x_15, x_84, x_17, x_141); if (lean_obj_tag(x_144) == 0) { lean_object* x_145; lean_object* x_146; @@ -23405,7 +23460,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_156 = l_Lean_Elab_Term_elabSubst___lambda__8(x_151, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_155, x_12, x_13, x_14, x_15, x_84, x_17, x_152); +x_156 = l_Lean_Elab_Term_elabSubst___lambda__9(x_151, x_9, x_19, x_94, x_10, x_1, x_87, x_8, x_6, x_155, x_12, x_13, x_14, x_15, x_84, x_17, x_152); if (lean_obj_tag(x_156) == 0) { lean_object* x_157; lean_object* x_158; @@ -23644,7 +23699,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_185 = l_Lean_Elab_Term_elabSubst___lambda__8(x_179, x_9, x_19, x_171, x_10, x_1, x_87, x_8, x_6, x_184, x_12, x_13, x_14, x_15, x_84, x_17, x_180); +x_185 = l_Lean_Elab_Term_elabSubst___lambda__9(x_179, x_9, x_19, x_171, x_10, x_1, x_87, x_8, x_6, x_184, x_12, x_13, x_14, x_15, x_84, x_17, x_180); if (lean_obj_tag(x_185) == 0) { lean_object* x_186; lean_object* x_187; @@ -23904,7 +23959,7 @@ lean_inc(x_8); lean_inc(x_1); lean_inc(x_10); lean_inc(x_9); -x_212 = l_Lean_Elab_Term_elabSubst___lambda__8(x_206, x_9, x_19, x_171, x_10, x_1, x_87, x_8, x_6, x_211, x_12, x_13, x_14, x_15, x_84, x_17, x_207); +x_212 = l_Lean_Elab_Term_elabSubst___lambda__9(x_206, x_9, x_19, x_171, x_10, x_1, x_87, x_8, x_6, x_211, x_12, x_13, x_14, x_15, x_84, x_17, x_207); if (lean_obj_tag(x_212) == 0) { lean_object* x_213; lean_object* x_214; @@ -24247,7 +24302,7 @@ return x_42; else { lean_object* x_43; lean_object* x_44; -x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__6), 12, 4); +x_43 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__7), 12, 4); lean_closure_set(x_43, 0, x_23); lean_closure_set(x_43, 1, x_2); lean_closure_set(x_43, 2, x_3); @@ -24337,61 +24392,6 @@ return x_68; } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_14 = l_Lean_Meta_mkEqSymm(x_3, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -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(0); -x_18 = lean_apply_12(x_1, x_2, x_15, x_5, x_4, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_16); -return x_18; -} -else -{ -uint8_t x_19; -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_19 = !lean_is_exclusive(x_14); -if (x_19 == 0) -{ -return x_14; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_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; -} -} -} -} static lean_object* _init_l_Lean_Elab_Term_elabSubst___closed__1() { _start: { @@ -24399,7 +24399,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__3; x_2 = l_Lean_Elab_Term_elabAnonymousCtor___lambda__3___closed__8; x_3 = l___regBuiltin_Lean_Elab_Term_elabCoe___closed__5; -x_4 = l_Lean_Elab_Term_elabSubst___lambda__5___closed__3; +x_4 = l_Lean_Elab_Term_elabSubst___lambda__6___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -24493,7 +24493,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabSubst___closed__12() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\nbut its left hand side is not mentioned in the type", 52); +x_1 = lean_mk_string_from_bytes("\nbut neither side of the equality is mentioned in the type", 58); return x_1; } } @@ -24756,7 +24756,7 @@ x_55 = lean_ctor_get(x_54, 1); lean_inc(x_55); if (lean_obj_tag(x_11) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_dec(x_17); lean_dec(x_1); x_56 = lean_ctor_get(x_41, 1); @@ -24767,6 +24767,8 @@ lean_inc(x_57); lean_dec(x_54); x_58 = lean_ctor_get(x_55, 0); lean_inc(x_58); +x_59 = lean_ctor_get(x_55, 1); +lean_inc(x_59); lean_dec(x_55); lean_inc(x_8); lean_inc(x_7); @@ -24774,120 +24776,147 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_59 = l_Lean_Elab_Term_elabTerm(x_23, x_28, x_24, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_56); -if (lean_obj_tag(x_59) == 0) +x_60 = l_Lean_Elab_Term_elabTerm(x_23, x_28, x_24, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_56); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_60, 0); lean_inc(x_61); -lean_dec(x_59); +x_62 = lean_ctor_get(x_60, 1); +lean_inc(x_62); +lean_dec(x_60); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_60); -x_62 = lean_infer_type(x_60, x_5, x_6, x_7, x_8, x_61); -if (lean_obj_tag(x_62) == 0) +lean_inc(x_61); +x_63 = lean_infer_type(x_61, x_5, x_6, x_7, x_8, x_62); +if (lean_obj_tag(x_63) == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_62, 1); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_64 = lean_ctor_get(x_63, 0); lean_inc(x_64); -lean_dec(x_62); -x_65 = lean_box(0); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_58); -lean_inc(x_63); -x_66 = l_Lean_Meta_kabstract(x_63, x_58, x_65, x_5, x_6, x_7, x_8, x_64); -if (lean_obj_tag(x_66) == 0) +lean_inc(x_64); +x_67 = l_Lean_Meta_kabstract(x_64, x_58, x_66, x_5, x_6, x_7, x_8, x_65); +if (lean_obj_tag(x_67) == 0) { -lean_object* x_67; lean_object* x_68; uint8_t x_69; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); -lean_dec(x_66); -x_69 = l_Lean_Expr_hasLooseBVars(x_67); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); lean_dec(x_67); -lean_dec(x_60); -lean_dec(x_58); +lean_inc(x_61); +lean_inc(x_57); +x_70 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__4___boxed), 14, 2); +lean_closure_set(x_70, 0, x_57); +lean_closure_set(x_70, 1, x_61); +x_71 = l_Lean_Expr_hasLooseBVars(x_68); +if (x_71 == 0) +{ +lean_object* x_72; +lean_dec(x_68); +lean_dec(x_61); lean_dec(x_57); -x_70 = l_Lean_indentExpr(x_33); -x_71 = l_Lean_Elab_Term_elabSubst___closed__9; -x_72 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_70); -x_73 = l_Lean_Elab_Term_elabSubst___closed__11; -x_74 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -x_75 = l_Lean_indentExpr(x_39); -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_74); -lean_ctor_set(x_76, 1, x_75); -x_77 = l_Lean_Elab_Term_elabSubst___closed__13; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_59); +lean_inc(x_64); +x_72 = l_Lean_Meta_kabstract(x_64, x_59, x_66, x_5, x_6, x_7, x_8, x_69); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; uint8_t x_75; +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 = l_Lean_Expr_hasLooseBVars(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +lean_dec(x_73); +lean_dec(x_70); +lean_dec(x_59); +lean_dec(x_58); +x_76 = l_Lean_indentExpr(x_33); +x_77 = l_Lean_Elab_Term_elabSubst___closed__9; x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_indentExpr(x_63); +lean_ctor_set(x_78, 0, x_77); +lean_ctor_set(x_78, 1, x_76); +x_79 = l_Lean_Elab_Term_elabSubst___closed__11; x_80 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_80, 0, x_78); lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_Elab_Term_elabCoeFunNotation___closed__4; +x_81 = l_Lean_indentExpr(x_39); x_82 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_82, x_3, x_4, x_5, x_6, x_7, x_8, x_68); +x_83 = l_Lean_Elab_Term_elabSubst___closed__13; +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +x_85 = l_Lean_indentExpr(x_64); +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_Elab_Term_elabCoeFunNotation___closed__4; +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_74); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_84 = !lean_is_exclusive(x_83); -if (x_84 == 0) +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) { -return x_83; +return x_89; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_83, 0); -x_86 = lean_ctor_get(x_83, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_83); -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_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_89, 0); +x_92 = lean_ctor_get(x_89, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_89); +x_93 = lean_alloc_ctor(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_88; lean_object* x_89; -lean_dec(x_63); +lean_object* x_94; lean_object* x_95; +lean_dec(x_64); lean_dec(x_39); -x_88 = lean_box(0); -x_89 = l_Lean_Elab_Term_elabSubst___lambda__4(x_58, x_67, x_57, x_60, x_33, x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_68); -return x_89; +x_94 = lean_box(0); +x_95 = l_Lean_Elab_Term_elabSubst___lambda__5(x_70, x_73, x_33, x_58, x_59, x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_74); +return x_95; } } else { -uint8_t x_90; -lean_dec(x_63); -lean_dec(x_60); +uint8_t x_96; +lean_dec(x_70); +lean_dec(x_64); +lean_dec(x_59); lean_dec(x_58); -lean_dec(x_57); lean_dec(x_39); lean_dec(x_33); lean_dec(x_8); @@ -24896,30 +24925,44 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_90 = !lean_is_exclusive(x_66); -if (x_90 == 0) +x_96 = !lean_is_exclusive(x_72); +if (x_96 == 0) { -return x_66; +return x_72; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_66, 0); -x_92 = lean_ctor_get(x_66, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_66); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_72, 0); +x_98 = lean_ctor_get(x_72, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_72); +x_99 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +return x_99; } } } else { -uint8_t x_94; -lean_dec(x_60); +lean_object* x_100; lean_object* x_101; +lean_dec(x_70); +lean_dec(x_64); +lean_dec(x_39); +x_100 = lean_box(0); +x_101 = l_Lean_Elab_Term_elabSubst___lambda__4(x_57, x_61, x_68, x_33, x_58, x_59, x_100, x_3, x_4, x_5, x_6, x_7, x_8, x_69); +lean_dec(x_59); +return x_101; +} +} +else +{ +uint8_t x_102; +lean_dec(x_64); +lean_dec(x_61); +lean_dec(x_59); lean_dec(x_58); lean_dec(x_57); lean_dec(x_39); @@ -24930,29 +24973,31 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_94 = !lean_is_exclusive(x_62); -if (x_94 == 0) +x_102 = !lean_is_exclusive(x_67); +if (x_102 == 0) { -return x_62; +return x_67; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_62, 0); -x_96 = lean_ctor_get(x_62, 1); -lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_62); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_103 = lean_ctor_get(x_67, 0); +x_104 = lean_ctor_get(x_67, 1); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_67); +x_105 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 1, x_104); +return x_105; } } } else { -uint8_t x_98; +uint8_t x_106; +lean_dec(x_61); +lean_dec(x_59); lean_dec(x_58); lean_dec(x_57); lean_dec(x_39); @@ -24963,77 +25008,111 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_98 = !lean_is_exclusive(x_59); -if (x_98 == 0) +x_106 = !lean_is_exclusive(x_63); +if (x_106 == 0) { -return x_59; +return x_63; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_59, 0); -x_100 = lean_ctor_get(x_59, 1); -lean_inc(x_100); -lean_inc(x_99); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_63, 0); +x_108 = lean_ctor_get(x_63, 1); +lean_inc(x_108); +lean_inc(x_107); +lean_dec(x_63); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; +} +} +} +else +{ +uint8_t x_110; lean_dec(x_59); -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; +lean_dec(x_58); +lean_dec(x_57); +lean_dec(x_39); +lean_dec(x_33); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_110 = !lean_is_exclusive(x_60); +if (x_110 == 0) +{ +return x_60; +} +else +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; +x_111 = lean_ctor_get(x_60, 0); +x_112 = lean_ctor_get(x_60, 1); +lean_inc(x_112); +lean_inc(x_111); +lean_dec(x_60); +x_113 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_113, 0, x_111); +lean_ctor_set(x_113, 1, x_112); +return x_113; } } } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_102 = lean_ctor_get(x_41, 1); -lean_inc(x_102); +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_114 = lean_ctor_get(x_41, 1); +lean_inc(x_114); lean_dec(x_41); -x_103 = lean_ctor_get(x_54, 0); -lean_inc(x_103); +x_115 = lean_ctor_get(x_54, 0); +lean_inc(x_115); lean_dec(x_54); -x_104 = lean_ctor_get(x_55, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_55, 1); -lean_inc(x_105); +x_116 = lean_ctor_get(x_55, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_55, 1); +lean_inc(x_117); lean_dec(x_55); -x_106 = lean_ctor_get(x_11, 0); -lean_inc(x_106); +x_118 = lean_ctor_get(x_11, 0); +lean_inc(x_118); lean_dec(x_11); -x_107 = lean_box(0); +x_119 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_105); -lean_inc(x_106); -x_108 = l_Lean_Meta_kabstract(x_106, x_105, x_107, x_5, x_6, x_7, x_8, x_102); -if (lean_obj_tag(x_108) == 0) +lean_inc(x_117); +lean_inc(x_118); +x_120 = l_Lean_Meta_kabstract(x_118, x_117, x_119, x_5, x_6, x_7, x_8, x_114); +if (lean_obj_tag(x_120) == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; -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_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; +x_121 = lean_ctor_get(x_120, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_120, 1); +lean_inc(x_122); +lean_dec(x_120); lean_inc(x_17); lean_inc(x_23); lean_inc(x_1); -lean_inc(x_106); -lean_inc(x_103); -x_111 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__9___boxed), 18, 6); -lean_closure_set(x_111, 0, x_103); -lean_closure_set(x_111, 1, x_106); -lean_closure_set(x_111, 2, x_1); -lean_closure_set(x_111, 3, x_23); -lean_closure_set(x_111, 4, x_17); -lean_closure_set(x_111, 5, x_28); -x_112 = l_Lean_Expr_hasLooseBVars(x_109); -if (x_112 == 0) -{ -lean_object* x_113; -lean_dec(x_109); -lean_dec(x_103); +lean_inc(x_118); +lean_inc(x_115); +x_123 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__10___boxed), 18, 6); +lean_closure_set(x_123, 0, x_115); +lean_closure_set(x_123, 1, x_118); +lean_closure_set(x_123, 2, x_1); +lean_closure_set(x_123, 3, x_23); +lean_closure_set(x_123, 4, x_17); +lean_closure_set(x_123, 5, x_28); +x_124 = l_Lean_Expr_hasLooseBVars(x_121); +if (x_124 == 0) +{ +lean_object* x_125; +lean_dec(x_121); +lean_dec(x_115); lean_dec(x_23); lean_dec(x_17); lean_dec(x_1); @@ -25041,92 +25120,92 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_104); -lean_inc(x_106); -x_113 = l_Lean_Meta_kabstract(x_106, x_104, x_107, x_5, x_6, x_7, x_8, x_110); -if (lean_obj_tag(x_113) == 0) +lean_inc(x_116); +lean_inc(x_118); +x_125 = l_Lean_Meta_kabstract(x_118, x_116, x_119, x_5, x_6, x_7, x_8, x_122); +if (lean_obj_tag(x_125) == 0) { -lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_116 = l_Lean_Expr_hasLooseBVars(x_114); -if (x_116 == 0) +lean_object* x_126; lean_object* x_127; uint8_t x_128; +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +x_128 = l_Lean_Expr_hasLooseBVars(x_126); +if (x_128 == 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; lean_object* x_129; lean_object* x_130; uint8_t x_131; -lean_dec(x_114); -lean_dec(x_111); -lean_dec(x_105); -lean_dec(x_104); -x_117 = l_Lean_indentExpr(x_106); -x_118 = l_Lean_Elab_Term_elabSubst___closed__15; -x_119 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_118); -lean_ctor_set(x_119, 1, x_117); -x_120 = l_Lean_Elab_Term_elabSubst___closed__17; -x_121 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -x_122 = l_Lean_indentExpr(x_33); -x_123 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -x_124 = l_Lean_Elab_Term_elabSubst___closed__19; -x_125 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -x_126 = l_Lean_indentExpr(x_39); -x_127 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_127, 0, x_125); -lean_ctor_set(x_127, 1, x_126); -x_128 = l_Lean_Elab_Term_elabSubst___closed__21; -x_129 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -x_130 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_129, x_3, x_4, x_5, x_6, x_7, x_8, x_115); +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; +lean_dec(x_126); +lean_dec(x_123); +lean_dec(x_117); +lean_dec(x_116); +x_129 = l_Lean_indentExpr(x_118); +x_130 = l_Lean_Elab_Term_elabSubst___closed__15; +x_131 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_129); +x_132 = l_Lean_Elab_Term_elabSubst___closed__17; +x_133 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +x_134 = l_Lean_indentExpr(x_33); +x_135 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +x_136 = l_Lean_Elab_Term_elabSubst___closed__19; +x_137 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_138 = l_Lean_indentExpr(x_39); +x_139 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +x_140 = l_Lean_Elab_Term_elabSubst___closed__21; +x_141 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_141, 0, x_139); +lean_ctor_set(x_141, 1, x_140); +x_142 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_141, x_3, x_4, x_5, x_6, x_7, x_8, x_127); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_131 = !lean_is_exclusive(x_130); -if (x_131 == 0) +x_143 = !lean_is_exclusive(x_142); +if (x_143 == 0) { -return x_130; +return x_142; } else { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_130, 0); -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -lean_inc(x_132); -lean_dec(x_130); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -return x_134; +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_142, 0); +x_145 = lean_ctor_get(x_142, 1); +lean_inc(x_145); +lean_inc(x_144); +lean_dec(x_142); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +return x_146; } } else { -lean_object* x_135; lean_object* x_136; -lean_dec(x_106); +lean_object* x_147; lean_object* x_148; +lean_dec(x_118); lean_dec(x_39); -x_135 = lean_box(0); -x_136 = l_Lean_Elab_Term_elabSubst___lambda__10(x_111, x_114, x_33, x_104, x_105, x_135, x_3, x_4, x_5, x_6, x_7, x_8, x_115); -return x_136; +x_147 = lean_box(0); +x_148 = l_Lean_Elab_Term_elabSubst___lambda__5(x_123, x_126, x_33, x_116, x_117, x_147, x_3, x_4, x_5, x_6, x_7, x_8, x_127); +return x_148; } } else { -uint8_t x_137; -lean_dec(x_111); -lean_dec(x_106); -lean_dec(x_105); -lean_dec(x_104); +uint8_t x_149; +lean_dec(x_123); +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); lean_dec(x_39); lean_dec(x_33); lean_dec(x_8); @@ -25135,43 +25214,43 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_137 = !lean_is_exclusive(x_113); -if (x_137 == 0) +x_149 = !lean_is_exclusive(x_125); +if (x_149 == 0) { -return x_113; +return x_125; } else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_138 = lean_ctor_get(x_113, 0); -x_139 = lean_ctor_get(x_113, 1); -lean_inc(x_139); -lean_inc(x_138); -lean_dec(x_113); -x_140 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_140, 0, x_138); -lean_ctor_set(x_140, 1, x_139); -return x_140; +lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_150 = lean_ctor_get(x_125, 0); +x_151 = lean_ctor_get(x_125, 1); +lean_inc(x_151); +lean_inc(x_150); +lean_dec(x_125); +x_152 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_152, 0, x_150); +lean_ctor_set(x_152, 1, x_151); +return x_152; } } } else { -lean_object* x_141; lean_object* x_142; -lean_dec(x_111); +lean_object* x_153; lean_object* x_154; +lean_dec(x_123); lean_dec(x_39); -x_141 = lean_box(0); -x_142 = l_Lean_Elab_Term_elabSubst___lambda__9(x_103, x_106, x_1, x_23, x_17, x_28, x_109, x_33, x_104, x_105, x_141, x_3, x_4, x_5, x_6, x_7, x_8, x_110); -return x_142; +x_153 = lean_box(0); +x_154 = l_Lean_Elab_Term_elabSubst___lambda__10(x_115, x_118, x_1, x_23, x_17, x_28, x_121, x_33, x_116, x_117, x_153, x_3, x_4, x_5, x_6, x_7, x_8, x_122); +return x_154; } } else { -uint8_t x_143; -lean_dec(x_106); -lean_dec(x_105); -lean_dec(x_104); -lean_dec(x_103); +uint8_t x_155; +lean_dec(x_118); +lean_dec(x_117); +lean_dec(x_116); +lean_dec(x_115); lean_dec(x_39); lean_dec(x_33); lean_dec(x_23); @@ -25183,23 +25262,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_143 = !lean_is_exclusive(x_108); -if (x_143 == 0) +x_155 = !lean_is_exclusive(x_120); +if (x_155 == 0) { -return x_108; +return x_120; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_108, 0); -x_145 = lean_ctor_get(x_108, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_108); -x_146 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; +lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_156 = lean_ctor_get(x_120, 0); +x_157 = lean_ctor_get(x_120, 1); +lean_inc(x_157); +lean_inc(x_156); +lean_dec(x_120); +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; } } } @@ -25207,7 +25286,7 @@ return x_146; } else { -uint8_t x_147; +uint8_t x_159; lean_dec(x_39); lean_dec(x_33); lean_dec(x_23); @@ -25220,29 +25299,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_147 = !lean_is_exclusive(x_41); -if (x_147 == 0) +x_159 = !lean_is_exclusive(x_41); +if (x_159 == 0) { return x_41; } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_148 = lean_ctor_get(x_41, 0); -x_149 = lean_ctor_get(x_41, 1); -lean_inc(x_149); -lean_inc(x_148); +lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_41, 0); +x_161 = lean_ctor_get(x_41, 1); +lean_inc(x_161); +lean_inc(x_160); lean_dec(x_41); -x_150 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -return x_150; +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; } } } else { -uint8_t x_151; +uint8_t x_163; lean_dec(x_33); lean_dec(x_23); lean_dec(x_17); @@ -25254,29 +25333,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_151 = !lean_is_exclusive(x_35); -if (x_151 == 0) +x_163 = !lean_is_exclusive(x_35); +if (x_163 == 0) { return x_35; } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_152 = lean_ctor_get(x_35, 0); -x_153 = lean_ctor_get(x_35, 1); -lean_inc(x_153); -lean_inc(x_152); +lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_164 = lean_ctor_get(x_35, 0); +x_165 = lean_ctor_get(x_35, 1); +lean_inc(x_165); +lean_inc(x_164); lean_dec(x_35); -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; +x_166 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_166, 0, x_164); +lean_ctor_set(x_166, 1, x_165); +return x_166; } } } else { -uint8_t x_155; +uint8_t x_167; lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25287,29 +25366,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_155 = !lean_is_exclusive(x_32); -if (x_155 == 0) +x_167 = !lean_is_exclusive(x_32); +if (x_167 == 0) { return x_32; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_32, 0); -x_157 = lean_ctor_get(x_32, 1); -lean_inc(x_157); -lean_inc(x_156); +lean_object* x_168; lean_object* x_169; lean_object* x_170; +x_168 = lean_ctor_get(x_32, 0); +x_169 = lean_ctor_get(x_32, 1); +lean_inc(x_169); +lean_inc(x_168); lean_dec(x_32); -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; +x_170 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_170, 0, x_168); +lean_ctor_set(x_170, 1, x_169); +return x_170; } } } else { -uint8_t x_159; +uint8_t x_171; lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25320,23 +25399,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_159 = !lean_is_exclusive(x_26); -if (x_159 == 0) +x_171 = !lean_is_exclusive(x_26); +if (x_171 == 0) { return x_26; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_26, 0); -x_161 = lean_ctor_get(x_26, 1); -lean_inc(x_161); -lean_inc(x_160); +lean_object* x_172; lean_object* x_173; lean_object* x_174; +x_172 = lean_ctor_get(x_26, 0); +x_173 = lean_ctor_get(x_26, 1); +lean_inc(x_173); +lean_inc(x_172); lean_dec(x_26); -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; +x_174 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_174, 0, x_172); +lean_ctor_set(x_174, 1, x_173); +return x_174; } } } @@ -25344,7 +25423,7 @@ return x_162; } else { -uint8_t x_163; +uint8_t x_175; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25352,23 +25431,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_163 = !lean_is_exclusive(x_10); -if (x_163 == 0) +x_175 = !lean_is_exclusive(x_10); +if (x_175 == 0) { return x_10; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_164 = lean_ctor_get(x_10, 0); -x_165 = lean_ctor_get(x_10, 1); -lean_inc(x_165); -lean_inc(x_164); +lean_object* x_176; lean_object* x_177; lean_object* x_178; +x_176 = lean_ctor_get(x_10, 0); +x_177 = lean_ctor_get(x_10, 1); +lean_inc(x_177); +lean_inc(x_176); lean_dec(x_10); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_164); -lean_ctor_set(x_166, 1, x_165); -return x_166; +x_178 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_178, 0, x_176); +lean_ctor_set(x_178, 1, x_177); +return x_178; } } } @@ -25399,25 +25478,35 @@ lean_dec(x_4); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; +x_15 = l_Lean_Elab_Term_elabSubst___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_7); +lean_dec(x_6); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___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: { lean_object* x_14; -x_14 = l_Lean_Elab_Term_elabSubst___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_14 = l_Lean_Elab_Term_elabSubst___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); lean_dec(x_6); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_15; -x_15 = l_Lean_Elab_Term_elabSubst___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, x_14); +x_15 = l_Lean_Elab_Term_elabSubst___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, x_14); lean_dec(x_7); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__8___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -25438,11 +25527,11 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Lean_Elab_Term_elabSubst___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, x_14, x_15, x_16, x_17); +x_18 = l_Lean_Elab_Term_elabSubst___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__9___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__10___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -25464,20 +25553,11 @@ lean_object* x_18 = _args[17]; _start: { lean_object* x_19; -x_19 = l_Lean_Elab_Term_elabSubst___lambda__9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +x_19 = l_Lean_Elab_Term_elabSubst___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, x_13, x_14, x_15, x_16, x_17, x_18); lean_dec(x_11); return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSubst___lambda__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: -{ -lean_object* x_14; -x_14 = l_Lean_Elab_Term_elabSubst___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, x_13); -lean_dec(x_6); -return x_14; -} -} static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst___closed__1() { _start: { @@ -25534,7 +25614,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSubst_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(418u); +x_1 = lean_unsigned_to_nat(422u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26025,7 +26105,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(420u); +x_1 = lean_unsigned_to_nat(424u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26037,7 +26117,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(429u); +x_1 = lean_unsigned_to_nat(433u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26065,7 +26145,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(420u); +x_1 = lean_unsigned_to_nat(424u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26077,7 +26157,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStateRefT_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(420u); +x_1 = lean_unsigned_to_nat(424u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26255,7 +26335,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(431u); +x_1 = lean_unsigned_to_nat(435u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26267,7 +26347,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(433u); +x_1 = lean_unsigned_to_nat(437u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26295,7 +26375,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(431u); +x_1 = lean_unsigned_to_nat(435u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26307,7 +26387,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoindex_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(431u); +x_1 = lean_unsigned_to_nat(435u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -26467,7 +26547,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__6() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__3; x_2 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(443u); +x_3 = lean_unsigned_to_nat(447u); x_4 = lean_unsigned_to_nat(54u); x_5 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__5; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -26480,7 +26560,7 @@ static lean_object* _init_l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__7() lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__3; x_2 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__4; -x_3 = lean_unsigned_to_nat(444u); +x_3 = lean_unsigned_to_nat(448u); x_4 = lean_unsigned_to_nat(55u); x_5 = l_Lean_Elab_Term_elabUnsafe___lambda__1___closed__5; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -27178,7 +27258,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabUnsafe_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(436u); +x_1 = lean_unsigned_to_nat(440u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27190,7 +27270,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabUnsafe_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(456u); +x_1 = lean_unsigned_to_nat(460u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27218,7 +27298,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabUnsafe_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(436u); +x_1 = lean_unsigned_to_nat(440u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -27230,7 +27310,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabUnsafe_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(436u); +x_1 = lean_unsigned_to_nat(440u); x_2 = lean_unsigned_to_nat(14u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28155,7 +28235,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRunElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(459u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28167,7 +28247,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRunElab_declRange___ _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(476u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28195,7 +28275,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRunElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(459u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28207,7 +28287,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRunElab_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(459u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28878,7 +28958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHaveI_declRange___cl _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(478u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28890,7 +28970,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHaveI_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(484u); +x_1 = lean_unsigned_to_nat(488u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28918,7 +28998,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHaveI_declRange___cl _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(478u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -28930,7 +29010,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHaveI_declRange___cl _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(478u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29290,7 +29370,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetI_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(486u); +x_1 = lean_unsigned_to_nat(490u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29302,7 +29382,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetI_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(496u); +x_1 = lean_unsigned_to_nat(500u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29330,7 +29410,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetI_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(486u); +x_1 = lean_unsigned_to_nat(490u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -29342,7 +29422,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetI_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(486u); +x_1 = lean_unsigned_to_nat(490u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -30565,18 +30645,18 @@ l_Lean_Elab_Term_elabSubst___lambda__4___closed__1 = _init_l_Lean_Elab_Term_elab lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__4___closed__1); l_Lean_Elab_Term_elabSubst___lambda__4___closed__2 = _init_l_Lean_Elab_Term_elabSubst___lambda__4___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__4___closed__2); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__1 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__1); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__2 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__2); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__3 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__3); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__4 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__4); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__5 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__5); -l_Lean_Elab_Term_elabSubst___lambda__5___closed__6 = _init_l_Lean_Elab_Term_elabSubst___lambda__5___closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__5___closed__6); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__1 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__1); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__2 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__2); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__3 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__3); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__4 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__4); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__5 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__5); +l_Lean_Elab_Term_elabSubst___lambda__6___closed__6 = _init_l_Lean_Elab_Term_elabSubst___lambda__6___closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_elabSubst___lambda__6___closed__6); l_Lean_Elab_Term_elabSubst___closed__1 = _init_l_Lean_Elab_Term_elabSubst___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_elabSubst___closed__1); l_Lean_Elab_Term_elabSubst___closed__2 = _init_l_Lean_Elab_Term_elabSubst___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index 0b135d6d5476..24d0ab77c467 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.BuiltinTerm -// Imports: Lean.Elab.Open Lean.Elab.SetOption Lean.Elab.Eval +// Imports: Lean.Meta.Diagnostics Lean.Elab.Open Lean.Elab.SetOption Lean.Elab.Eval #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -74,7 +74,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v lean_object* lean_uint32_to_nat(uint32_t); static lean_object* l___regBuiltin_Lean_Elab_Term_elabRawNatLit___closed__3; lean_object* l_System_FilePath_join(lean_object*, lean_object*); -extern lean_object* l_Lean_diag; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__4; lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Elab_Term_elabClear___spec__19___boxed(lean_object*, lean_object*); @@ -132,6 +131,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeOf___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabHole___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__1; static lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Term_elabOpen___spec__34___lambda__1___closed__1; +static lean_object* l_Lean_Elab_Term_elabSetOption___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMVar_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabRawNatLit___closed__2; lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_Elab_Term_addDotCompletionInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -298,6 +298,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabIncludeStr_declRange( LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabStrLit___closed__2; static lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_evalFilePathUnsafe___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabQuotedName___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_elabSyntheticHole___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabBadCDot___closed__1; @@ -307,6 +308,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWithAnnotateTerm___lambda__1___box static lean_object* l___regBuiltin_Lean_Elab_Term_elabQuotedName___closed__2; LEAN_EXPORT lean_object* l_Lean_popScope___at_Lean_Elab_Term_elabOpen___spec__45___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_mkFreshTypeMVarFor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf___closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_OpenDecl_elabOpenDecl___spec__6(size_t, size_t, lean_object*); @@ -314,6 +316,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabCompletion___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabIncludeStr_declRange___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSyntheticHole_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabBadCDot_declRange___closed__2; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabOmission_declRange___closed__7; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_elabOpen___spec__35(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -339,6 +342,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabWaitIfContainsMVar_declRan LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabOpen___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isDelayedAssigned___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Meta_reportDiag(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_ResolveName_0__Lean_resolveGlobalConstCore___at_Lean_Elab_Term_elabOpen___spec__14(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_elabCharLit_declRange___closed__1; @@ -456,6 +460,7 @@ static lean_object* l_Lean_Elab_Term_elabOpen___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm_declRange(lean_object*); 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___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Term_elabOpen___spec__34___closed__1; +static lean_object* l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabPipeCompletion(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabQuotedName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -600,6 +605,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDeclName___lambda__1___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabWaitIfTypeMVar(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_mkFreshExprMVarImpl(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at_Lean_Elab_Term_elabSyntheticHole___spec__4(lean_object*); +extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_elabOpen___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen___spec__30___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); @@ -902,7 +908,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange___closed_ static lean_object* l___regBuiltin_Lean_Elab_Term_elabStrLit___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_elabClear___spec__8___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_BuiltinTerm_0__Lean_Elab_Term_evalFilePathUnsafe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___boxed(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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRange___closed__1; static lean_object* l_Lean_Elab_Term_elabNumLit___lambda__1___closed__1; @@ -934,6 +939,7 @@ lean_object* l_Lean_MetavarContext_findUserName_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMVar_declRange(lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabOpen_declRange___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSort___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_declRange___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabClear___spec__4(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*); @@ -1204,7 +1210,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProp_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(14u); +x_1 = lean_unsigned_to_nat(15u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1216,7 +1222,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProp_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(15u); +x_1 = lean_unsigned_to_nat(16u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1244,7 +1250,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProp_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(14u); +x_1 = lean_unsigned_to_nat(15u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1256,7 +1262,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabProp_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(14u); +x_1 = lean_unsigned_to_nat(15u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1480,7 +1486,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSort_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1492,7 +1498,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSort_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(24u); +x_1 = lean_unsigned_to_nat(25u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1520,7 +1526,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSort_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1532,7 +1538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSort_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1719,7 +1725,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); +x_1 = lean_unsigned_to_nat(27u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1731,7 +1737,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(27u); +x_1 = lean_unsigned_to_nat(28u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1759,7 +1765,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); +x_1 = lean_unsigned_to_nat(27u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1771,7 +1777,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeStx_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); +x_1 = lean_unsigned_to_nat(27u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2093,7 +2099,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(35u); +x_1 = lean_unsigned_to_nat(36u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2105,7 +2111,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(39u); +x_1 = lean_unsigned_to_nat(40u); x_2 = lean_unsigned_to_nat(78u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2133,7 +2139,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(35u); +x_1 = lean_unsigned_to_nat(36u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2145,7 +2151,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabPipeCompletion_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(35u); +x_1 = lean_unsigned_to_nat(36u); x_2 = lean_unsigned_to_nat(60u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2462,7 +2468,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCompletion_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(41u); +x_1 = lean_unsigned_to_nat(42u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2474,7 +2480,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCompletion_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(56u); +x_1 = lean_unsigned_to_nat(57u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2502,7 +2508,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCompletion_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(41u); +x_1 = lean_unsigned_to_nat(42u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2514,7 +2520,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCompletion_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(41u); +x_1 = lean_unsigned_to_nat(42u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2760,7 +2766,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHole_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(58u); +x_1 = lean_unsigned_to_nat(59u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2772,7 +2778,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHole_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(62u); +x_1 = lean_unsigned_to_nat(63u); x_2 = lean_unsigned_to_nat(11u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2800,7 +2806,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHole_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(58u); +x_1 = lean_unsigned_to_nat(59u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2812,7 +2818,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabHole_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(58u); +x_1 = lean_unsigned_to_nat(59u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4639,7 +4645,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSyntheticHole_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(64u); +x_1 = lean_unsigned_to_nat(65u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4651,7 +4657,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSyntheticHole_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(100u); +x_1 = lean_unsigned_to_nat(101u); x_2 = lean_unsigned_to_nat(97u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4679,7 +4685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSyntheticHole_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(64u); +x_1 = lean_unsigned_to_nat(65u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4691,7 +4697,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSyntheticHole_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(64u); +x_1 = lean_unsigned_to_nat(65u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4844,7 +4850,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOmission_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4856,7 +4862,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOmission_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(108u); +x_1 = lean_unsigned_to_nat(109u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4884,7 +4890,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOmission_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4896,7 +4902,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOmission_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(102u); +x_1 = lean_unsigned_to_nat(103u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5329,7 +5335,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5341,7 +5347,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(121u); +x_1 = lean_unsigned_to_nat(122u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5368,7 +5374,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5380,7 +5386,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetMVar_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(110u); +x_1 = lean_unsigned_to_nat(111u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5744,7 +5750,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeMVar_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5756,7 +5762,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeMVar_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5784,7 +5790,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeMVar_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5796,7 +5802,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeMVar_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(129u); +x_1 = lean_unsigned_to_nat(130u); x_2 = lean_unsigned_to_nat(60u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6102,7 +6108,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMV _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6114,7 +6120,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMV _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(142u); +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); @@ -6142,7 +6148,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMV _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6154,7 +6160,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfTypeContainsMV _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(76u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6392,7 +6398,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfContainsMVar_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(144u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6404,7 +6410,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfContainsMVar_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(150u); +x_1 = lean_unsigned_to_nat(151u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6432,7 +6438,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfContainsMVar_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(144u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6444,7 +6450,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWaitIfContainsMVar_d _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(144u); +x_1 = lean_unsigned_to_nat(145u); x_2 = lean_unsigned_to_nat(68u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6717,7 +6723,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabByTactic_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(159u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6729,7 +6735,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabByTactic_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(164u); +x_1 = lean_unsigned_to_nat(165u); x_2 = lean_unsigned_to_nat(75u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6757,7 +6763,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabByTactic_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(159u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6769,7 +6775,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabByTactic_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(159u); +x_1 = lean_unsigned_to_nat(160u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6932,7 +6938,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(167u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6944,7 +6950,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); +x_1 = lean_unsigned_to_nat(168u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6972,7 +6978,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(167u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6984,7 +6990,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNoImplicitLambda_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(166u); +x_1 = lean_unsigned_to_nat(167u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7193,7 +7199,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabBadCDot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(170u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7205,7 +7211,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabBadCDot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(173u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(104u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7233,7 +7239,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabBadCDot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(170u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7245,7 +7251,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabBadCDot_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(170u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7435,7 +7441,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___c _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(176u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7447,7 +7453,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___c _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(179u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7475,7 +7481,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___c _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(176u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7487,7 +7493,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabStrLit_declRange___c _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(176u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8066,7 +8072,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange___c _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(188u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8078,7 +8084,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange___c _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(197u); x_2 = lean_unsigned_to_nat(10u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8106,7 +8112,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange___c _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(188u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8118,7 +8124,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabNumLit_declRange___c _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(188u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8270,7 +8276,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRawNatLit_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(198u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8282,7 +8288,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRawNatLit_declRange_ _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(202u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8310,7 +8316,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRawNatLit_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(198u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8322,7 +8328,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabRawNatLit_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(198u); +x_1 = lean_unsigned_to_nat(199u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8796,7 +8802,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8808,7 +8814,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRa _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(214u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8836,7 +8842,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -8848,7 +8854,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(204u); +x_1 = lean_unsigned_to_nat(205u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9030,7 +9036,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCharLit_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(216u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9042,7 +9048,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCharLit_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(219u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9070,7 +9076,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCharLit_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(216u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9082,7 +9088,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabCharLit_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(216u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9236,7 +9242,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabQuotedName_declRange _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(221u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9248,7 +9254,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabQuotedName_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(224u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9276,7 +9282,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabQuotedName_declRange _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(221u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9288,7 +9294,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabQuotedName_declRange _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(221u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9471,7 +9477,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDoubleQuotedName_dec _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(226u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9483,7 +9489,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDoubleQuotedName_dec _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(227u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9511,7 +9517,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDoubleQuotedName_dec _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(226u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9523,7 +9529,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDoubleQuotedName_dec _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(226u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9814,7 +9820,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDeclName_declRange__ _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(229u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9826,7 +9832,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDeclName_declRange__ _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(232u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9854,7 +9860,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDeclName_declRange__ _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(229u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -9866,7 +9872,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDeclName_declRange__ _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(229u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10027,7 +10033,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithDeclName_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(234u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10039,7 +10045,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithDeclName_declRan _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(238u); x_2 = lean_unsigned_to_nat(73u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10067,7 +10073,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithDeclName_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(234u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10079,7 +10085,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithDeclName_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(234u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10251,7 +10257,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange___c _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(240u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10263,7 +10269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange___c _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(241u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10291,7 +10297,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange___c _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(240u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10303,7 +10309,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabTypeOf_declRange___c _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(240u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10862,7 +10868,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10874,7 +10880,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(277u); +x_1 = lean_unsigned_to_nat(278u); x_2 = lean_unsigned_to_nat(97u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10902,7 +10908,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -10914,7 +10920,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureTypeOf_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(270u); +x_1 = lean_unsigned_to_nat(271u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11083,7 +11089,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_d _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(280u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11095,7 +11101,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_d _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(283u); x_2 = lean_unsigned_to_nat(82u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11123,7 +11129,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_d _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(280u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -11135,7 +11141,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabEnsureExpectedType_d _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(280u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14455,7 +14461,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabClear_declRange___cl _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(285u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14467,7 +14473,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabClear_declRange___cl _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(303u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14495,7 +14501,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabClear_declRange___cl _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(285u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14507,7 +14513,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabClear_declRange___cl _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(285u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21559,7 +21565,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOpen_declRange___clo _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(305u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21571,7 +21577,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOpen_declRange___clo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(312u); +x_1 = lean_unsigned_to_nat(313u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21599,7 +21605,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOpen_declRange___clo _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(305u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21611,7 +21617,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabOpen_declRange___clo _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(305u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22041,7 +22047,7 @@ return x_71; } } } -static lean_object* _init_l_Lean_Elab_Term_elabSetOption___closed__1() { +static lean_object* _init_l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -22049,11 +22055,451 @@ x_1 = l_Lean_maxRecDepth; return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___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, uint8_t x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +uint8_t x_14; +x_14 = !lean_is_exclusive(x_11); +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_20; +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = lean_ctor_get(x_11, 2); +lean_dec(x_16); +x_17 = l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1; +x_18 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_17); +lean_ctor_set(x_11, 4, x_18); +lean_ctor_set(x_11, 2, x_1); +lean_ctor_set_uint8(x_11, sizeof(void*)*11 + 1, x_2); +x_19 = 1; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_8); +lean_inc(x_7); +x_20 = l_Lean_Elab_Term_elabTerm(x_3, x_4, x_19, x_19, x_5, x_6, x_7, x_8, x_11, x_12, x_13); +if (lean_obj_tag(x_20) == 0) +{ +if (x_9 == 0) +{ +uint8_t x_21; +lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +return x_20; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = l_Lean_Meta_reportDiag(x_7, x_8, x_11, x_12, x_26); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_25); +return x_27; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +else +{ +uint8_t x_32; +lean_dec(x_25); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) +{ +return x_27; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +else +{ +if (x_9 == 0) +{ +uint8_t x_36; +lean_dec(x_11); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +x_36 = !lean_is_exclusive(x_20); +if (x_36 == 0) +{ +return x_20; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_20, 0); +x_38 = lean_ctor_get(x_20, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_20); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_20, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_20, 1); +lean_inc(x_41); +lean_dec(x_20); +x_42 = l_Lean_Meta_reportDiag(x_7, x_8, x_11, x_12, x_41); +if (lean_obj_tag(x_42) == 0) +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +lean_ctor_set_tag(x_42, 1); +lean_ctor_set(x_42, 0, x_40); +return x_42; +} +else +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_40); +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_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; +x_51 = lean_ctor_get(x_11, 0); +x_52 = lean_ctor_get(x_11, 1); +x_53 = lean_ctor_get(x_11, 3); +x_54 = lean_ctor_get(x_11, 5); +x_55 = lean_ctor_get(x_11, 6); +x_56 = lean_ctor_get(x_11, 7); +x_57 = lean_ctor_get(x_11, 8); +x_58 = lean_ctor_get(x_11, 9); +x_59 = lean_ctor_get(x_11, 10); +x_60 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); +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_inc(x_52); +lean_inc(x_51); +lean_dec(x_11); +x_61 = l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1; +x_62 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_61); +x_63 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_63, 0, x_51); +lean_ctor_set(x_63, 1, x_52); +lean_ctor_set(x_63, 2, x_1); +lean_ctor_set(x_63, 3, x_53); +lean_ctor_set(x_63, 4, x_62); +lean_ctor_set(x_63, 5, x_54); +lean_ctor_set(x_63, 6, x_55); +lean_ctor_set(x_63, 7, x_56); +lean_ctor_set(x_63, 8, x_57); +lean_ctor_set(x_63, 9, x_58); +lean_ctor_set(x_63, 10, x_59); +lean_ctor_set_uint8(x_63, sizeof(void*)*11, x_60); +lean_ctor_set_uint8(x_63, sizeof(void*)*11 + 1, x_2); +x_64 = 1; +lean_inc(x_12); +lean_inc(x_63); +lean_inc(x_8); +lean_inc(x_7); +x_65 = l_Lean_Elab_Term_elabTerm(x_3, x_4, x_64, x_64, x_5, x_6, x_7, x_8, x_63, x_12, x_13); +if (lean_obj_tag(x_65) == 0) +{ +if (x_9 == 0) +{ +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +lean_dec(x_63); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +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; +x_70 = lean_ctor_get(x_65, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_65, 1); +lean_inc(x_71); +lean_dec(x_65); +x_72 = l_Lean_Meta_reportDiag(x_7, x_8, x_63, x_12, x_71); +if (lean_obj_tag(x_72) == 0) +{ +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_ctor_get(x_72, 1); +lean_inc(x_73); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_74 = x_72; +} else { + lean_dec_ref(x_72); + x_74 = lean_box(0); +} +if (lean_is_scalar(x_74)) { + x_75 = lean_alloc_ctor(0, 2, 0); +} else { + x_75 = x_74; +} +lean_ctor_set(x_75, 0, x_70); +lean_ctor_set(x_75, 1, x_73); +return x_75; +} +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_70); +x_76 = lean_ctor_get(x_72, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_72, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_78 = x_72; +} else { + lean_dec_ref(x_72); + x_78 = lean_box(0); +} +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); +} else { + x_79 = x_78; +} +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +} +else +{ +if (x_9 == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +lean_dec(x_63); +lean_dec(x_12); +lean_dec(x_8); +lean_dec(x_7); +x_80 = lean_ctor_get(x_65, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_65, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_82 = x_65; +} else { + lean_dec_ref(x_65); + x_82 = lean_box(0); +} +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); +} else { + 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; lean_object* x_86; +x_84 = lean_ctor_get(x_65, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_65, 1); +lean_inc(x_85); +lean_dec(x_65); +x_86 = l_Lean_Meta_reportDiag(x_7, x_8, x_63, x_12, x_85); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_86, 1); +lean_inc(x_87); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_88 = x_86; +} else { + lean_dec_ref(x_86); + x_88 = lean_box(0); +} +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); +} else { + x_89 = x_88; + lean_ctor_set_tag(x_89, 1); +} +lean_ctor_set(x_89, 0, x_84); +lean_ctor_set(x_89, 1, x_87); +return x_89; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +lean_dec(x_84); +x_90 = lean_ctor_get(x_86, 0); +lean_inc(x_90); +x_91 = lean_ctor_get(x_86, 1); +lean_inc(x_91); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_92 = x_86; +} else { + lean_dec_ref(x_86); + x_92 = lean_box(0); +} +if (lean_is_scalar(x_92)) { + x_93 = lean_alloc_ctor(1, 2, 0); +} else { + x_93 = x_92; +} +lean_ctor_set(x_93, 0, x_90); +lean_ctor_set(x_93, 1, x_91); +return x_93; +} +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Term_elabSetOption___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("diagnostics", 11); +return x_1; +} +} static lean_object* _init_l_Lean_Elab_Term_elabSetOption___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Term_elabSetOption___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Term_elabSetOption___closed__3() { +_start: +{ lean_object* x_1; -x_1 = l_Lean_diag; +x_1 = l_Lean_diagnostics; return x_1; } } @@ -22071,10 +22517,11 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); +lean_inc(x_11); x_14 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1(x_11, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -22082,74 +22529,164 @@ lean_inc(x_16); lean_dec(x_14); x_17 = lean_unsigned_to_nat(5u); x_18 = l_Lean_Syntax_getArg(x_1, x_17); -x_19 = !lean_is_exclusive(x_7); -if (x_19 == 0) +lean_dec(x_1); +x_19 = l_Lean_Syntax_getId(x_11); +lean_dec(x_11); +x_20 = l_Lean_Elab_Term_elabSetOption___closed__2; +x_21 = lean_name_eq(x_19, x_20); +lean_dec(x_19); +x_22 = l_Lean_Elab_Term_elabSetOption___closed__3; +x_23 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_22); +x_24 = lean_st_ref_get(x_8, x_16); +if (x_21 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; lean_object* x_27; -x_20 = lean_ctor_get(x_7, 4); -lean_dec(x_20); -x_21 = lean_ctor_get(x_7, 2); -lean_dec(x_21); -x_22 = l_Lean_Elab_Term_elabSetOption___closed__1; -x_23 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_15, x_22); -x_24 = l_Lean_Elab_Term_elabSetOption___closed__2; -x_25 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_24); -lean_ctor_set(x_7, 4, x_23); -lean_ctor_set(x_7, 2, x_15); -lean_ctor_set_uint8(x_7, sizeof(void*)*11 + 1, x_25); -x_26 = 1; -x_27 = l_Lean_Elab_Term_elabTerm(x_18, x_2, x_26, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_16); -return x_27; +lean_object* x_64; lean_object* x_65; uint8_t x_66; +x_64 = lean_ctor_get(x_24, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_24, 1); +lean_inc(x_65); +lean_dec(x_24); +x_66 = 0; +x_25 = x_66; +x_26 = x_64; +x_27 = x_65; +goto block_63; } 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; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; -x_28 = lean_ctor_get(x_7, 0); -x_29 = lean_ctor_get(x_7, 1); -x_30 = lean_ctor_get(x_7, 3); -x_31 = lean_ctor_get(x_7, 5); -x_32 = lean_ctor_get(x_7, 6); -x_33 = lean_ctor_get(x_7, 7); -x_34 = lean_ctor_get(x_7, 8); -x_35 = lean_ctor_get(x_7, 9); -x_36 = lean_ctor_get(x_7, 10); -x_37 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); -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_object* x_67; lean_object* x_68; uint8_t x_69; +x_67 = lean_ctor_get(x_24, 0); +lean_inc(x_67); +x_68 = lean_ctor_get(x_24, 1); +lean_inc(x_68); +lean_dec(x_24); +x_69 = 1; +x_25 = x_69; +x_26 = x_67; +x_27 = x_68; +goto block_63; +} +block_63: +{ +lean_object* x_28; uint8_t x_29; uint8_t x_30; +x_28 = lean_ctor_get(x_26, 0); lean_inc(x_28); -lean_dec(x_7); -x_38 = l_Lean_Elab_Term_elabSetOption___closed__1; -x_39 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_15, x_38); -x_40 = l_Lean_Elab_Term_elabSetOption___closed__2; -x_41 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_15, x_40); -x_42 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_42, 0, x_28); -lean_ctor_set(x_42, 1, x_29); -lean_ctor_set(x_42, 2, x_15); -lean_ctor_set(x_42, 3, x_30); -lean_ctor_set(x_42, 4, x_39); -lean_ctor_set(x_42, 5, x_31); -lean_ctor_set(x_42, 6, x_32); -lean_ctor_set(x_42, 7, x_33); -lean_ctor_set(x_42, 8, x_34); -lean_ctor_set(x_42, 9, x_35); -lean_ctor_set(x_42, 10, x_36); -lean_ctor_set_uint8(x_42, sizeof(void*)*11, x_37); -lean_ctor_set_uint8(x_42, sizeof(void*)*11 + 1, x_41); -x_43 = 1; -x_44 = l_Lean_Elab_Term_elabTerm(x_18, x_2, x_43, x_43, x_3, x_4, x_5, x_6, x_42, x_8, x_16); -return x_44; +lean_dec(x_26); +x_29 = l_Lean_Kernel_isDiagnosticsEnabled(x_28); +lean_dec(x_28); +if (x_29 == 0) +{ +if (x_23 == 0) +{ +uint8_t x_59; +x_59 = 1; +x_30 = x_59; +goto block_58; +} +else +{ +uint8_t x_60; +x_60 = 0; +x_30 = x_60; +goto block_58; } } else { -uint8_t x_45; +if (x_23 == 0) +{ +uint8_t x_61; +x_61 = 0; +x_30 = x_61; +goto block_58; +} +else +{ +uint8_t x_62; +x_62 = 1; +x_30 = x_62; +goto block_58; +} +} +block_58: +{ +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_st_ref_take(x_8, x_27); +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; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_35 = lean_ctor_get(x_32, 0); +x_36 = lean_ctor_get(x_32, 4); +lean_dec(x_36); +x_37 = l_Lean_Kernel_enableDiag(x_35, x_23); +x_38 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen___spec__2___closed__4; +lean_ctor_set(x_32, 4, x_38); +lean_ctor_set(x_32, 0, x_37); +x_39 = lean_st_ref_set(x_8, x_32, x_33); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_box(0); +x_42 = l_Lean_Elab_Term_elabSetOption___lambda__1(x_15, x_23, x_18, x_2, x_3, x_4, x_5, x_6, x_25, x_41, x_7, x_8, 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; +x_43 = lean_ctor_get(x_32, 0); +x_44 = lean_ctor_get(x_32, 1); +x_45 = lean_ctor_get(x_32, 2); +x_46 = lean_ctor_get(x_32, 3); +x_47 = lean_ctor_get(x_32, 5); +x_48 = lean_ctor_get(x_32, 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_32); +x_49 = l_Lean_Kernel_enableDiag(x_43, x_23); +x_50 = l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabOpen___spec__2___closed__4; +x_51 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_44); +lean_ctor_set(x_51, 2, x_45); +lean_ctor_set(x_51, 3, x_46); +lean_ctor_set(x_51, 4, x_50); +lean_ctor_set(x_51, 5, x_47); +lean_ctor_set(x_51, 6, x_48); +x_52 = lean_st_ref_set(x_8, x_51, x_33); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +x_54 = lean_box(0); +x_55 = l_Lean_Elab_Term_elabSetOption___lambda__1(x_15, x_23, x_18, x_2, x_3, x_4, x_5, x_6, x_25, x_54, x_7, x_8, x_53); +return x_55; +} +} +else +{ +lean_object* x_56; lean_object* x_57; +x_56 = lean_box(0); +x_57 = l_Lean_Elab_Term_elabSetOption___lambda__1(x_15, x_23, x_18, x_2, x_3, x_4, x_5, x_6, x_25, x_56, x_7, x_8, x_27); +return x_57; +} +} +} +} +else +{ +uint8_t x_70; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -22157,23 +22694,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_45 = !lean_is_exclusive(x_14); -if (x_45 == 0) +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_14); +if (x_70 == 0) { return x_14; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_14, 0); -x_47 = lean_ctor_get(x_14, 1); -lean_inc(x_47); -lean_inc(x_46); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_14, 0); +x_72 = lean_ctor_get(x_14, 1); +lean_inc(x_72); +lean_inc(x_71); lean_dec(x_14); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +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; } } } @@ -22217,13 +22755,17 @@ lean_dec(x_5); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabSetOption___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_10; -x_10 = l_Lean_Elab_Term_elabSetOption(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_1); -return x_10; +uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_14 = lean_unbox(x_2); +lean_dec(x_2); +x_15 = lean_unbox(x_9); +lean_dec(x_9); +x_16 = l_Lean_Elab_Term_elabSetOption___lambda__1(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_15, x_10, x_11, x_12, x_13); +lean_dec(x_10); +return x_16; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__1() { @@ -22270,7 +22812,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSetOption___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSetOption), 9, 0); return x_1; } } @@ -22290,7 +22832,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange_ _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(315u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22302,8 +22844,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(317u); -x_2 = lean_unsigned_to_nat(33u); +x_1 = lean_unsigned_to_nat(322u); +x_2 = lean_unsigned_to_nat(18u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -22317,7 +22859,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange___closed__1; x_2 = lean_unsigned_to_nat(34u); x_3 = l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange___closed__2; -x_4 = lean_unsigned_to_nat(33u); +x_4 = lean_unsigned_to_nat(18u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -22330,7 +22872,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange_ _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(315u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22342,7 +22884,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabSetOption_declRange_ _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(315u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22522,7 +23064,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(319u); +x_1 = lean_unsigned_to_nat(324u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22534,7 +23076,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm_dec _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(328u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22562,7 +23104,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(319u); +x_1 = lean_unsigned_to_nat(324u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22574,7 +23116,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabWithAnnotateTerm_dec _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(319u); +x_1 = lean_unsigned_to_nat(324u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22944,7 +23486,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIncludeStr_declRange _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(336u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22956,7 +23498,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIncludeStr_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(345u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22984,7 +23526,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIncludeStr_declRange _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(336u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22996,7 +23538,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabIncludeStr_declRange _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(336u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23042,6 +23584,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } +lean_object* initialize_Lean_Meta_Diagnostics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Open(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_SetOption(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Eval(uint8_t builtin, lean_object*); @@ -23050,6 +23593,9 @@ LEAN_EXPORT lean_object* initialize_Lean_Elab_BuiltinTerm(uint8_t builtin, lean_ lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Lean_Meta_Diagnostics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Elab_Open(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -24207,10 +24753,14 @@ l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__ lean_mark_persistent(l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__3); l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__4 = _init_l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__4(); lean_mark_persistent(l_Lean_Elab_elabSetOption___at_Lean_Elab_Term_elabSetOption___spec__1___closed__4); +l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1 = _init_l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_elabSetOption___lambda__1___closed__1); l_Lean_Elab_Term_elabSetOption___closed__1 = _init_l_Lean_Elab_Term_elabSetOption___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_elabSetOption___closed__1); l_Lean_Elab_Term_elabSetOption___closed__2 = _init_l_Lean_Elab_Term_elabSetOption___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_elabSetOption___closed__2); +l_Lean_Elab_Term_elabSetOption___closed__3 = _init_l_Lean_Elab_Term_elabSetOption___closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_elabSetOption___closed__3); l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__1 = _init_l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__1); l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__2 = _init_l___regBuiltin_Lean_Elab_Term_elabSetOption___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index ef5abcb32f99..b868b75811e9 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Command -// Imports: Lean.Elab.Binders Lean.Elab.SyntheticMVars Lean.Elab.SetOption Lean.Language.Basic +// Imports: Lean.Meta.Diagnostics Lean.Elab.Binders Lean.Elab.SyntheticMVars Lean.Elab.SetOption Lean.Language.Basic #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -13,19 +13,20 @@ #ifdef __cplusplus extern "C" { #endif -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_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(lean_object*, size_t, size_t, lean_object*); 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*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_instMonadCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_liftAttrM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__1; @@ -35,16 +36,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandD static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__27; 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2974____closed__1; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqOfDecidableEq___rarg(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; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10; +extern lean_object* l_Lean_diagExt; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13; extern lean_object* l_Lean_profiler; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); @@ -52,21 +51,19 @@ uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___spec__1(lean_object*, size_t, size_t, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; 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*, 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*); 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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__23; @@ -76,7 +73,7 @@ static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__6; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___closed__2; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__12; lean_object* lean_private_to_user_name(lean_object*); -extern lean_object* l_Lean_diag; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___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_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,6 +85,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__L static lean_object* l_Lean_liftCommandElabM___rarg___closed__4; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__15; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__3; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); @@ -103,6 +101,7 @@ static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabComma extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); @@ -112,10 +111,12 @@ lean_object* l_Lean_Elab_mkMessageCore(lean_object*, lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore(lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_NameHashSet_insert___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_traceState___default; static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__2; @@ -129,17 +130,19 @@ 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_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__20; 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*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3(lean_object*, size_t, size_t, 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_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabCommand___spec__18(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); -static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___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__3___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLevel___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_expandDeclId___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -148,8 +151,10 @@ static lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___closed__2; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__5; 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_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13(lean_object*); lean_object* l_String_removeLeadingSpaces(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*); @@ -168,15 +173,14 @@ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_applyVisibility___at_Lean_Elab_Command_expandDeclId___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_mkState(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__3; lean_object* l_Lean_Syntax_getArgs(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__21; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftTermElabM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -187,21 +191,24 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEq; static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withSetOptionIn___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM(lean_object*); lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2974____closed__2; LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg(lean_object*, lean_object*); -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_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__3; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; +lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___closed__2; @@ -211,12 +218,16 @@ lean_object* lean_environment_find(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, 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__12___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__3; static lean_object* l_Lean_Elab_Command_elabCommand___closed__5; @@ -227,7 +238,6 @@ LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec_ 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_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3346____closed__3; extern lean_object* l_Lean_LocalContext_empty; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -236,7 +246,6 @@ 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*); @@ -246,20 +255,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAd uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_instMonad(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withSetOptionIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___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_Elab_Command_instMonadLogCommandElabM___lambda__2___boxed(lean_object*, lean_object*, 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_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___closed__1; -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_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*); -static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___closed__1; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); @@ -274,26 +280,27 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef___boxed(lean_object*, lean_o static lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___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_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___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___at_Lean_Elab_Command_runLinters___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__3; static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__2___closed__1; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__18___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___lambda__1(lean_object*, lean_object*); static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__1; uint8_t l_Lean_isReservedName(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_InternalExceptionId_getName(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg___boxed(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(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withSetOptionIn___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___at_Lean_Elab_Command_elabCommand___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -301,7 +308,7 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_log___at_Lean_Elab_Command_runLinters___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -312,19 +319,20 @@ lean_object* l_Lean_Meta_reportDiag(lean_object*, lean_object*, lean_object*, le LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Command_elabCommand___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__3; size_t lean_usize_of_nat(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__17(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3346_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___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__10___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___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*, 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__4___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_Elab_Command_instMonadRecDepthCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908_(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_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___closed__2; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__4; @@ -353,6 +361,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg___boxed(lean_object*, static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getBracketedBinderIds(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); @@ -361,14 +370,15 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_el LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_nextMacroScope___default; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__9(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__4; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__25; 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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__3; @@ -385,6 +395,7 @@ 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; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM(lean_object*); @@ -394,22 +405,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabComma 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_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLiftTIOCommandElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); 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_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_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_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* l_Lean_Syntax_getKind(lean_object*); static lean_object* l_Lean_withSetOptionIn___closed__1; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__11; @@ -422,19 +430,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Context_macroStack___default; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); extern lean_object* l___private_Lean_DocString_0__Lean_docStringExt; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___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_EXPORT lean_object* l_Lean_liftCommandElabM(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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__22; 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_object* l_Lean_Meta_MetaM_run___rarg(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_547____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___boxed(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; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__10; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_Command_instMonadEnvCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__2; @@ -445,11 +453,8 @@ 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_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7(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*); 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*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17; @@ -467,16 +472,16 @@ LEAN_EXPORT lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommand___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isStrLit_x3f(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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8___boxed(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_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*); 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; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_Elab_Command_elabCommand___closed__4; @@ -485,9 +490,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_modifyScope(lean_object*, lean_obje lean_object* l_Lean_Elab_InfoTree_substitute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__6(lean_object*, 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_Command_initFn____x40_Lean_Elab_Command___hyg_2538_(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__18; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); @@ -495,20 +500,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lea 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; -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*, 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_547____closed__9; lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__7; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15___rarg___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_Elab_Command_expandDeclId___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__10(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2538____closed__1; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___lambda__1___closed__5; static lean_object* l_Lean_Elab_Command_instInhabitedScope___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12___lambda__1(lean_object*, lean_object*); -static double l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); @@ -516,13 +518,14 @@ lean_object* l_Lean_extractMacroScopes(lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__3; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__12___rarg(lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__7(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__14___boxed(lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__14; @@ -535,7 +538,6 @@ 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*); 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_403____closed__26; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -543,6 +545,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId static lean_object* l_Lean_Elab_Command_Scope_varDecls___default___closed__1; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__5; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__1; +extern lean_object* l_Lean_trace_profiler_threshold; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_liftCommandElabM___rarg___closed__1; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -553,15 +556,14 @@ static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed_ static lean_object* l_Lean_liftCommandElabM___rarg___closed__2; 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; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_currNamespace___default; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2538____closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__2; +extern lean_object* l_Lean_diagnostics; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14; static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__1; @@ -575,10 +577,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScope___rarg(lean_object*, lean_ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_ioErrorToMessage(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__1; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___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_Lean_Elab_Command_liftTermElabM___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_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__2(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_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__2; 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*); @@ -586,9 +589,9 @@ 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*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Command_runLinters___spec__12(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__3; +static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; 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*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__16; @@ -624,29 +627,31 @@ lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__5; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Command___hyg_403_; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3346____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__19; +static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__1; static lean_object* l_Lean_Elab_Command_State_messages___default___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_elabCommand___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_515_(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__3; @@ -654,18 +659,17 @@ static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__ static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2; 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*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(lean_object*, size_t, size_t, lean_object*); 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_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*, lean_object*, uint8_t, 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__12___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_openDecls___default; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__3; static lean_object* l_Lean_Elab_Command_instAddMessageContextCommandElabM___closed__1; @@ -687,18 +691,21 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermE 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*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__18(lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_addLinter___closed__1; lean_object* l_ReaderT_instFunctorOfMonad___rarg(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_withFreshMacroScope(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2; static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__2; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__8; @@ -711,9 +718,11 @@ 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*); static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___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_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*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Option_get_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -725,11 +734,12 @@ static lean_object* l_Lean_Elab_Command_elabCommand___closed__6; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Command_modifyScope___spec__1(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftTermElabM___spec__14(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashable; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddErrorMessageContextCommandElabM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -744,6 +754,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Comm LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___rarg(lean_object*, lean_object*); uint8_t l_Lean_isStructure(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___closed__1; @@ -755,21 +766,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabComma static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__2; lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_DataValue_sameCtor(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_expandDeclId___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__28; 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_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_initFn____x40_Lean_Elab_Command___hyg_547____closed__5; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__14; +lean_object* l_Lean_Kernel_resetDiag___lambda__1(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessages___closed__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); @@ -783,6 +792,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId__ 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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3; static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__3; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__7; @@ -792,13 +802,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__ static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedScope; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__21(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); @@ -806,33 +814,32 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varUIds___default; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__6; 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*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__17; -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*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_Scope_isNoncomputable___default; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftTermElabM___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KVMap_findCore(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___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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_Command_getScope___rarg___boxed(lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler; uint8_t l_Lean_PersistentArray_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1; lean_object* lean_io_error_to_string(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__6; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -844,7 +851,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Comma LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2974_(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_liftCommandElabM___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___boxed(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); @@ -852,12 +858,10 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters_ extern lean_object* l_Lean_trace_profiler_output; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__4; static lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___closed__3; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(lean_object*, size_t, size_t, lean_object*); lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_liftCommandElabM___rarg___closed__3; -static lean_object* l_Lean_Elab_Command_liftCoreM___rarg___closed__1; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__3; static lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed__1; @@ -865,9 +869,9 @@ lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Command_elabCommand___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkInfoTree___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField___at_Lean_Elab_Command_expandDeclId___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_addUnivLevel(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Command_liftTermElabM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -886,6 +890,7 @@ static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__2 static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__5; static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12; lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__4; static lean_object* l_Lean_Elab_Command_modifyScope___closed__1; static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___closed__1; @@ -894,28 +899,25 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Command_elabCommandTopLe LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__3___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__13(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_expandDeclId___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -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_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; 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*); static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_infoState___default; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__16(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_runLinters___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__1___closed__1; 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_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469_(lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__4; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3346____closed__2; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___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___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___boxed(lean_object*, lean_object*); @@ -927,20 +929,21 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___closed__1; lean_object* l_Lean_MessageData_ofName(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3(lean_object*, lean_object*, lean_object*); 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_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___spec__15(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore(lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__11___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11___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*); double lean_float_sub(double, double); @@ -3334,71 +3337,6 @@ x_6 = l_Lean_Elab_Command_mkMessageAux(x_1, x_2, x_3, x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_diag; -return x_1; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; 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; -x_4 = lean_ctor_get(x_2, 2); -x_5 = lean_ctor_get(x_2, 4); -x_6 = l_Lean_Elab_Command_instInhabitedScope; -x_7 = l_List_head_x21___rarg(x_6, x_4); -x_8 = lean_ctor_get(x_1, 0); -x_9 = lean_ctor_get(x_1, 1); -x_10 = lean_ctor_get(x_1, 2); -x_11 = lean_ctor_get(x_1, 5); -x_12 = lean_ctor_get(x_1, 6); -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -x_14 = lean_ctor_get(x_7, 2); -lean_inc(x_14); -x_15 = lean_ctor_get(x_7, 3); -lean_inc(x_15); -lean_dec(x_7); -x_16 = l_Lean_Core_getMaxHeartbeats(x_13); -x_17 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___closed__1; -x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_13, x_17); -x_19 = 0; -lean_inc(x_11); -lean_inc(x_12); -lean_inc(x_5); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -x_20 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_20, 0, x_8); -lean_ctor_set(x_20, 1, x_9); -lean_ctor_set(x_20, 2, x_13); -lean_ctor_set(x_20, 3, x_10); -lean_ctor_set(x_20, 4, x_5); -lean_ctor_set(x_20, 5, x_12); -lean_ctor_set(x_20, 6, x_14); -lean_ctor_set(x_20, 7, x_15); -lean_ctor_set(x_20, 8, x_3); -lean_ctor_set(x_20, 9, x_16); -lean_ctor_set(x_20, 10, x_11); -lean_ctor_set_uint8(x_20, sizeof(void*)*11, x_19); -lean_ctor_set_uint8(x_20, sizeof(void*)*11 + 1, x_18); -return x_20; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_1, x_2, x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; -} -} LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1(lean_object* x_1) { _start: { @@ -5494,7 +5432,7 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -5509,7 +5447,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x x_6 = lean_array_uget(x_4, x_3); x_7 = lean_unsigned_to_nat(0u); x_8 = lean_array_uset(x_4, x_3, x_7); -x_9 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_6); +x_9 = l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(x_1, x_6); x_10 = 1; x_11 = lean_usize_add(x_3, x_10); x_12 = lean_array_uset(x_8, x_3, x_9); @@ -5519,7 +5457,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -5537,44 +5475,42 @@ x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = !lean_is_exclusive(x_6); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_1, 6); -x_12 = l_Lean_replaceRef(x_10, x_11); +x_11 = l_Lean_replaceRef(x_10, x_1); lean_dec(x_10); -lean_ctor_set(x_6, 0, x_12); -x_13 = 1; -x_14 = lean_usize_add(x_3, x_13); -x_15 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_14; -x_4 = x_15; +lean_ctor_set(x_6, 0, x_11); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_14 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_13; +x_4 = x_14; goto _start; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_17 = lean_ctor_get(x_6, 0); -x_18 = lean_ctor_get(x_6, 1); -lean_inc(x_18); +lean_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_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_6, 1); lean_inc(x_17); +lean_inc(x_16); lean_dec(x_6); -x_19 = lean_ctor_get(x_1, 6); -x_20 = l_Lean_replaceRef(x_17, x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_18); -x_22 = 1; -x_23 = lean_usize_add(x_3, x_22); -x_24 = lean_array_uset(x_8, x_3, x_21); -x_3 = x_23; -x_4 = x_24; +x_18 = l_Lean_replaceRef(x_16, x_1); +lean_dec(x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_22 = lean_array_uset(x_8, x_3, x_19); +x_3 = x_21; +x_4 = x_22; goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) @@ -5589,7 +5525,7 @@ x_5 = lean_array_get_size(x_4); x_6 = lean_usize_of_nat(x_5); lean_dec(x_5); x_7 = 0; -x_8 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(x_1, x_6, x_7, x_4); +x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3(x_1, x_6, x_7, x_4); lean_ctor_set(x_2, 0, x_8); return x_2; } @@ -5603,7 +5539,7 @@ x_10 = lean_array_get_size(x_9); x_11 = lean_usize_of_nat(x_10); lean_dec(x_10); x_12 = 0; -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(x_1, x_11, x_12, x_9); +x_13 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__3(x_1, x_11, x_12, x_9); x_14 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_14, 0, x_13); return x_14; @@ -5621,7 +5557,7 @@ x_17 = lean_array_get_size(x_16); x_18 = lean_usize_of_nat(x_17); lean_dec(x_17); x_19 = 0; -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_18, x_19, x_16); +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(x_1, x_18, x_19, x_16); lean_ctor_set(x_2, 0, x_20); return x_2; } @@ -5635,7 +5571,7 @@ x_22 = lean_array_get_size(x_21); x_23 = lean_usize_of_nat(x_22); lean_dec(x_22); x_24 = 0; -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_23, x_24, x_21); +x_25 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(x_1, x_23, x_24, x_21); x_26 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_26, 0, x_25); return x_26; @@ -5643,7 +5579,7 @@ return x_26; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -5661,44 +5597,42 @@ x_8 = lean_array_uset(x_4, x_3, x_7); x_9 = !lean_is_exclusive(x_6); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; size_t x_13; size_t x_14; lean_object* x_15; +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_1, 6); -x_12 = l_Lean_replaceRef(x_10, x_11); +x_11 = l_Lean_replaceRef(x_10, x_1); lean_dec(x_10); -lean_ctor_set(x_6, 0, x_12); -x_13 = 1; -x_14 = lean_usize_add(x_3, x_13); -x_15 = lean_array_uset(x_8, x_3, x_6); -x_3 = x_14; -x_4 = x_15; +lean_ctor_set(x_6, 0, x_11); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_14 = lean_array_uset(x_8, x_3, x_6); +x_3 = x_13; +x_4 = x_14; goto _start; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; size_t x_22; size_t x_23; lean_object* x_24; -x_17 = lean_ctor_get(x_6, 0); -x_18 = lean_ctor_get(x_6, 1); -lean_inc(x_18); +lean_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_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_6, 1); lean_inc(x_17); +lean_inc(x_16); lean_dec(x_6); -x_19 = lean_ctor_get(x_1, 6); -x_20 = l_Lean_replaceRef(x_17, x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_18); -x_22 = 1; -x_23 = lean_usize_add(x_3, x_22); -x_24 = lean_array_uset(x_8, x_3, x_21); -x_3 = x_23; -x_4 = x_24; +x_18 = l_Lean_replaceRef(x_16, x_1); +lean_dec(x_16); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_22 = lean_array_uset(x_8, x_3, x_19); +x_3 = x_21; +x_4 = x_22; goto _start; } } } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; @@ -5708,12 +5642,12 @@ if (x_3 == 0) lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; lean_object* x_10; x_4 = lean_ctor_get(x_2, 0); x_5 = lean_ctor_get(x_2, 1); -x_6 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_4); +x_6 = l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(x_1, x_4); x_7 = lean_array_get_size(x_5); x_8 = lean_usize_of_nat(x_7); lean_dec(x_7); x_9 = 0; -x_10 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_8, x_9, x_5); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(x_1, x_8, x_9, x_5); lean_ctor_set(x_2, 1, x_10); lean_ctor_set(x_2, 0, x_6); return x_2; @@ -5731,12 +5665,12 @@ lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_dec(x_2); -x_16 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_11); +x_16 = l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(x_1, x_11); x_17 = lean_array_get_size(x_12); x_18 = lean_usize_of_nat(x_17); lean_dec(x_17); x_19 = 0; -x_20 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_18, x_19, x_12); +x_20 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(x_1, x_18, x_19, x_12); x_21 = lean_alloc_ctor(0, 4, sizeof(size_t)*1); lean_ctor_set(x_21, 0, x_16); lean_ctor_set(x_21, 1, x_20); @@ -5747,7 +5681,87 @@ return x_21; } } } -static lean_object* _init_l_Lean_Elab_Command_liftCoreM___rarg___closed__1() { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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) { +_start: +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +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_ctor_get(x_5, 4); +lean_dec(x_9); +x_10 = lean_ctor_get(x_5, 2); +lean_dec(x_10); +x_11 = l_Lean_Elab_Command_mkState___closed__1; +x_12 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_11); +lean_ctor_set(x_5, 4, x_12); +lean_ctor_set(x_5, 2, x_1); +lean_ctor_set_uint8(x_5, sizeof(void*)*11 + 1, x_2); +x_13 = lean_apply_3(x_3, x_5, x_6, x_7); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_14 = lean_ctor_get(x_5, 0); +x_15 = lean_ctor_get(x_5, 1); +x_16 = lean_ctor_get(x_5, 3); +x_17 = lean_ctor_get(x_5, 5); +x_18 = lean_ctor_get(x_5, 6); +x_19 = lean_ctor_get(x_5, 7); +x_20 = lean_ctor_get(x_5, 8); +x_21 = lean_ctor_get(x_5, 9); +x_22 = lean_ctor_get(x_5, 10); +x_23 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); +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_24 = l_Lean_Elab_Command_mkState___closed__1; +x_25 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_24); +x_26 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_26, 0, x_14); +lean_ctor_set(x_26, 1, x_15); +lean_ctor_set(x_26, 2, x_1); +lean_ctor_set(x_26, 3, x_16); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set(x_26, 5, x_17); +lean_ctor_set(x_26, 6, x_18); +lean_ctor_set(x_26, 7, x_19); +lean_ctor_set(x_26, 8, x_20); +lean_ctor_set(x_26, 9, x_21); +lean_ctor_set(x_26, 10, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*11, x_23); +lean_ctor_set_uint8(x_26, sizeof(void*)*11 + 1, x_2); +x_27 = lean_apply_3(x_3, x_26, x_6, x_7); +return x_27; +} +} +} +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagExt; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Kernel_resetDiag___lambda__1), 1, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -5758,10 +5772,18 @@ lean_ctor_set(x_2, 1, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; uint8_t x_34; x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -5774,754 +5796,1176 @@ lean_inc(x_9); x_10 = lean_ctor_get(x_8, 1); lean_inc(x_10); lean_dec(x_8); -x_11 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkCoreContext(x_2, x_6, x_9); -x_12 = lean_ctor_get(x_6, 6); +x_11 = lean_ctor_get(x_6, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_6, 2); lean_inc(x_12); -x_13 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 3); lean_inc(x_13); -x_14 = lean_ctor_get(x_6, 5); +x_14 = lean_ctor_get(x_6, 4); lean_inc(x_14); -x_15 = lean_ctor_get(x_6, 7); +x_15 = lean_ctor_get(x_6, 5); lean_inc(x_15); +x_16 = lean_ctor_get(x_6, 6); +lean_inc(x_16); +x_17 = lean_ctor_get(x_6, 7); +lean_inc(x_17); lean_dec(x_6); -x_16 = !lean_is_exclusive(x_12); -if (x_16 == 0) +x_18 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; +x_19 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__2; +x_20 = l_Lean_EnvExtension_modifyState___rarg(x_18, x_11, x_19); +x_21 = l_Lean_Elab_Command_instInhabitedScope; +x_22 = l_List_head_x21___rarg(x_21, x_12); +lean_dec(x_12); +x_23 = lean_ctor_get(x_2, 0); +x_24 = lean_ctor_get(x_2, 1); +x_25 = lean_ctor_get(x_2, 2); +x_26 = lean_ctor_get(x_2, 5); +x_27 = lean_ctor_get(x_2, 6); +x_28 = lean_ctor_get(x_22, 1); +lean_inc(x_28); +x_29 = lean_ctor_get(x_22, 2); +lean_inc(x_29); +x_30 = lean_ctor_get(x_22, 3); +lean_inc(x_30); +lean_dec(x_22); +x_31 = l_Lean_Core_getMaxHeartbeats(x_28); +x_32 = 0; +lean_inc(x_26); +lean_inc(x_27); +lean_inc(x_25); +lean_inc(x_28); +lean_inc(x_24); +lean_inc(x_23); +x_33 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_33, 0, x_23); +lean_ctor_set(x_33, 1, x_24); +lean_ctor_set(x_33, 2, x_28); +lean_ctor_set(x_33, 3, x_25); +lean_ctor_set(x_33, 4, x_14); +lean_ctor_set(x_33, 5, x_27); +lean_ctor_set(x_33, 6, x_29); +lean_ctor_set(x_33, 7, x_30); +lean_ctor_set(x_33, 8, x_9); +lean_ctor_set(x_33, 9, x_31); +lean_ctor_set(x_33, 10, x_26); +lean_ctor_set_uint8(x_33, sizeof(void*)*11, x_32); +lean_ctor_set_uint8(x_33, sizeof(void*)*11 + 1, x_32); +x_34 = !lean_is_exclusive(x_16); +if (x_34 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_17 = lean_ctor_get(x_12, 1); -lean_dec(x_17); -x_18 = lean_ctor_get(x_12, 0); -lean_dec(x_18); -x_19 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_20 = l_Lean_Elab_Command_State_messages___default___closed__3; -lean_ctor_set(x_12, 1, x_20); -lean_ctor_set(x_12, 0, x_19); -x_21 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; -x_22 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_23 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_23, 0, x_13); -lean_ctor_set(x_23, 1, x_21); -lean_ctor_set(x_23, 2, x_14); -lean_ctor_set(x_23, 3, x_15); -lean_ctor_set(x_23, 4, x_22); -lean_ctor_set(x_23, 5, x_20); -lean_ctor_set(x_23, 6, x_12); -x_102 = lean_st_mk_ref(x_23, x_10); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; uint8_t x_112; +x_35 = lean_ctor_get(x_16, 1); +lean_dec(x_35); +x_36 = lean_ctor_get(x_16, 0); +lean_dec(x_36); +x_37 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_38 = l_Lean_Elab_Command_State_messages___default___closed__3; +lean_ctor_set(x_16, 1, x_38); +lean_ctor_set(x_16, 0, x_37); +x_39 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; +x_40 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_40, 0, x_20); +lean_ctor_set(x_40, 1, x_13); +lean_ctor_set(x_40, 2, x_15); +lean_ctor_set(x_40, 3, x_17); +lean_ctor_set(x_40, 4, x_39); +lean_ctor_set(x_40, 5, x_38); +lean_ctor_set(x_40, 6, x_16); +x_102 = lean_st_mk_ref(x_40, x_10); x_103 = lean_ctor_get(x_102, 0); lean_inc(x_103); x_104 = lean_ctor_get(x_102, 1); lean_inc(x_104); lean_dec(x_102); -lean_inc(x_103); -x_105 = lean_apply_3(x_1, x_11, x_103, x_104); -if (lean_obj_tag(x_105) == 0) -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); -lean_inc(x_107); -lean_dec(x_105); -x_108 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_108, 0, x_106); -x_109 = lean_st_ref_get(x_103, x_107); -lean_dec(x_103); -x_110 = lean_ctor_get(x_109, 0); +x_105 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_106 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_105); +x_107 = lean_st_ref_get(x_103, x_104); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_107, 1); +lean_inc(x_109); +lean_dec(x_107); +x_110 = lean_ctor_get(x_108, 0); lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); -lean_inc(x_111); -lean_dec(x_109); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_108); -lean_ctor_set(x_112, 1, x_110); -x_24 = x_112; -x_25 = x_111; -goto block_101; -} -else -{ -uint8_t x_113; -x_113 = !lean_is_exclusive(x_105); -if (x_113 == 0) +lean_dec(x_108); +x_111 = l_Lean_Kernel_isDiagnosticsEnabled(x_110); +lean_dec(x_110); +if (x_111 == 0) { -lean_object* x_114; lean_object* x_115; uint8_t x_116; -x_114 = lean_ctor_get(x_105, 0); -x_115 = lean_ctor_get(x_105, 1); -x_116 = l_Lean_Exception_isRuntime(x_114); -if (x_116 == 0) +if (x_106 == 0) { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -lean_free_object(x_105); -x_117 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_117, 0, x_114); -x_118 = lean_st_ref_get(x_103, x_115); -lean_dec(x_103); -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 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_117); -lean_ctor_set(x_121, 1, x_119); -x_24 = x_121; -x_25 = x_120; -goto block_101; +uint8_t x_169; +x_169 = 1; +x_112 = x_169; +goto block_168; } else { -lean_dec(x_103); -return x_105; +x_112 = x_32; +goto block_168; } } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_122 = lean_ctor_get(x_105, 0); -x_123 = lean_ctor_get(x_105, 1); -lean_inc(x_123); -lean_inc(x_122); -lean_dec(x_105); -x_124 = l_Lean_Exception_isRuntime(x_122); -if (x_124 == 0) -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_125 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_125, 0, x_122); -x_126 = lean_st_ref_get(x_103, x_123); -lean_dec(x_103); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_126, 1); -lean_inc(x_128); -lean_dec(x_126); -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_125); -lean_ctor_set(x_129, 1, x_127); -x_24 = x_129; -x_25 = x_128; -goto block_101; +if (x_106 == 0) +{ +x_112 = x_32; +goto block_168; } else { -lean_object* x_130; -lean_dec(x_103); -x_130 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_130, 0, x_122); -lean_ctor_set(x_130, 1, x_123); -return x_130; -} +uint8_t x_170; +x_170 = 1; +x_112 = x_170; +goto block_168; } } block_101: { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_26 = lean_ctor_get(x_24, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_24, 1); -lean_inc(x_27); -lean_dec(x_24); -x_28 = lean_st_ref_take(x_3, x_25); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_27, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_27, 2); -lean_inc(x_32); -x_33 = lean_ctor_get(x_27, 3); -lean_inc(x_33); -x_34 = lean_ctor_get(x_27, 5); -lean_inc(x_34); -x_35 = lean_ctor_get(x_27, 6); -lean_inc(x_35); -lean_dec(x_27); -x_36 = !lean_is_exclusive(x_29); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; -x_37 = lean_ctor_get(x_29, 1); -x_38 = lean_ctor_get(x_29, 6); -x_39 = lean_ctor_get(x_29, 7); -lean_dec(x_39); -x_40 = lean_ctor_get(x_29, 5); -lean_dec(x_40); -x_41 = lean_ctor_get(x_29, 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; uint8_t x_54; +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_42 = l_Lean_PersistentArray_append___rarg(x_37, x_34); -x_43 = !lean_is_exclusive(x_38); -if (x_43 == 0) -{ -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_44 = lean_ctor_get(x_38, 1); -x_45 = lean_ctor_get(x_35, 1); -lean_inc(x_45); -lean_dec(x_35); -x_46 = l_Lean_PersistentArray_append___rarg(x_44, x_45); -lean_ctor_set(x_38, 1, x_46); -x_47 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); -lean_ctor_set(x_29, 7, x_47); -lean_ctor_set(x_29, 5, x_32); -lean_ctor_set(x_29, 1, x_42); -lean_ctor_set(x_29, 0, x_31); -x_48 = lean_st_ref_set(x_3, x_29, x_30); -if (lean_obj_tag(x_26) == 0) +x_45 = lean_st_ref_take(x_3, x_42); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +lean_dec(x_45); +x_48 = lean_ctor_get(x_44, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_44, 1); +lean_inc(x_49); +x_50 = lean_ctor_get(x_44, 2); +lean_inc(x_50); +x_51 = lean_ctor_get(x_44, 3); +lean_inc(x_51); +x_52 = lean_ctor_get(x_44, 5); +lean_inc(x_52); +x_53 = lean_ctor_get(x_44, 6); +lean_inc(x_53); +lean_dec(x_44); +x_54 = !lean_is_exclusive(x_46); +if (x_54 == 0) { -uint8_t x_49; -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_55 = lean_ctor_get(x_46, 1); +x_56 = lean_ctor_get(x_46, 6); +x_57 = lean_ctor_get(x_46, 7); +lean_dec(x_57); +x_58 = lean_ctor_get(x_46, 5); +lean_dec(x_58); +x_59 = lean_ctor_get(x_46, 3); +lean_dec(x_59); +x_60 = lean_ctor_get(x_46, 0); +lean_dec(x_60); +x_61 = l_Lean_PersistentArray_append___rarg(x_55, x_52); +x_62 = !lean_is_exclusive(x_56); +if (x_62 == 0) { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_48, 0); -lean_dec(x_50); -x_51 = lean_ctor_get(x_26, 0); -lean_inc(x_51); -lean_dec(x_26); -lean_ctor_set_tag(x_48, 1); -lean_ctor_set(x_48, 0, x_51); -return x_48; +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_ctor_get(x_56, 1); +x_64 = lean_ctor_get(x_53, 1); +lean_inc(x_64); +lean_dec(x_53); +x_65 = l_Lean_PersistentArray_append___rarg(x_63, x_64); +lean_ctor_set(x_56, 1, x_65); +x_66 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_27, x_51); +lean_ctor_set(x_46, 7, x_66); +lean_ctor_set(x_46, 5, x_50); +lean_ctor_set(x_46, 3, x_49); +lean_ctor_set(x_46, 1, x_61); +lean_ctor_set(x_46, 0, x_48); +x_67 = lean_st_ref_set(x_3, x_46, x_47); +x_68 = !lean_is_exclusive(x_67); +if (x_68 == 0) +{ +lean_object* x_69; +x_69 = lean_ctor_get(x_67, 0); +lean_dec(x_69); +lean_ctor_set(x_67, 0, x_43); +return x_67; } 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_ctor_get(x_26, 0); -lean_inc(x_53); -lean_dec(x_26); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_52); -return x_54; +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_67, 1); +lean_inc(x_70); +lean_dec(x_67); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_43); +lean_ctor_set(x_71, 1, x_70); +return x_71; } } else { -uint8_t x_55; -x_55 = !lean_is_exclusive(x_48); -if (x_55 == 0) -{ -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_48, 0); +uint8_t 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_72 = lean_ctor_get_uint8(x_56, sizeof(void*)*2); +x_73 = lean_ctor_get(x_56, 0); +x_74 = lean_ctor_get(x_56, 1); +lean_inc(x_74); +lean_inc(x_73); lean_dec(x_56); -x_57 = lean_ctor_get(x_26, 0); -lean_inc(x_57); -lean_dec(x_26); -lean_ctor_set(x_48, 0, x_57); -return x_48; +x_75 = lean_ctor_get(x_53, 1); +lean_inc(x_75); +lean_dec(x_53); +x_76 = l_Lean_PersistentArray_append___rarg(x_74, x_75); +x_77 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_77, 0, x_73); +lean_ctor_set(x_77, 1, x_76); +lean_ctor_set_uint8(x_77, sizeof(void*)*2, x_72); +x_78 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_27, x_51); +lean_ctor_set(x_46, 7, x_78); +lean_ctor_set(x_46, 6, x_77); +lean_ctor_set(x_46, 5, x_50); +lean_ctor_set(x_46, 3, x_49); +lean_ctor_set(x_46, 1, x_61); +lean_ctor_set(x_46, 0, x_48); +x_79 = lean_st_ref_set(x_3, x_46, x_47); +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_79)) { + lean_ctor_release(x_79, 0); + lean_ctor_release(x_79, 1); + x_81 = x_79; +} else { + lean_dec_ref(x_79); + x_81 = lean_box(0); } -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_48, 1); -lean_inc(x_58); -lean_dec(x_48); -x_59 = lean_ctor_get(x_26, 0); -lean_inc(x_59); -lean_dec(x_26); -x_60 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_60, 1, x_58); -return x_60; +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(0, 2, 0); +} else { + x_82 = x_81; } +lean_ctor_set(x_82, 0, x_43); +lean_ctor_set(x_82, 1, x_80); +return x_82; } } else { -uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_61 = lean_ctor_get_uint8(x_38, sizeof(void*)*2); -x_62 = lean_ctor_get(x_38, 0); -x_63 = lean_ctor_get(x_38, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_38); -x_64 = lean_ctor_get(x_35, 1); -lean_inc(x_64); -lean_dec(x_35); -x_65 = l_Lean_PersistentArray_append___rarg(x_63, x_64); -x_66 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_66, 0, x_62); -lean_ctor_set(x_66, 1, x_65); -lean_ctor_set_uint8(x_66, sizeof(void*)*2, x_61); -x_67 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); -lean_ctor_set(x_29, 7, x_67); -lean_ctor_set(x_29, 6, x_66); -lean_ctor_set(x_29, 5, x_32); -lean_ctor_set(x_29, 1, x_42); -lean_ctor_set(x_29, 0, x_31); -x_68 = lean_st_ref_set(x_3, x_29, x_30); -if (lean_obj_tag(x_26) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_70 = x_68; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; lean_object* x_89; lean_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_83 = lean_ctor_get(x_46, 1); +x_84 = lean_ctor_get(x_46, 2); +x_85 = lean_ctor_get(x_46, 4); +x_86 = lean_ctor_get(x_46, 6); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_46); +x_87 = l_Lean_PersistentArray_append___rarg(x_83, x_52); +x_88 = lean_ctor_get_uint8(x_86, sizeof(void*)*2); +x_89 = lean_ctor_get(x_86, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_86, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_91 = x_86; } else { - lean_dec_ref(x_68); - x_70 = lean_box(0); + lean_dec_ref(x_86); + x_91 = lean_box(0); } -x_71 = lean_ctor_get(x_26, 0); -lean_inc(x_71); -lean_dec(x_26); -if (lean_is_scalar(x_70)) { - x_72 = lean_alloc_ctor(1, 2, 0); +x_92 = lean_ctor_get(x_53, 1); +lean_inc(x_92); +lean_dec(x_53); +x_93 = l_Lean_PersistentArray_append___rarg(x_90, x_92); +if (lean_is_scalar(x_91)) { + x_94 = lean_alloc_ctor(0, 2, 1); } else { - x_72 = x_70; - lean_ctor_set_tag(x_72, 1); + x_94 = x_91; } -lean_ctor_set(x_72, 0, x_71); -lean_ctor_set(x_72, 1, x_69); -return x_72; -} -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_73 = lean_ctor_get(x_68, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_74 = x_68; +lean_ctor_set(x_94, 0, x_89); +lean_ctor_set(x_94, 1, x_93); +lean_ctor_set_uint8(x_94, sizeof(void*)*2, x_88); +x_95 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_27, x_51); +x_96 = lean_alloc_ctor(0, 8, 0); +lean_ctor_set(x_96, 0, x_48); +lean_ctor_set(x_96, 1, x_87); +lean_ctor_set(x_96, 2, x_84); +lean_ctor_set(x_96, 3, x_49); +lean_ctor_set(x_96, 4, x_85); +lean_ctor_set(x_96, 5, x_50); +lean_ctor_set(x_96, 6, x_94); +lean_ctor_set(x_96, 7, x_95); +x_97 = lean_st_ref_set(x_3, x_96, x_47); +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_68); - x_74 = lean_box(0); + lean_dec_ref(x_97); + x_99 = lean_box(0); } -x_75 = lean_ctor_get(x_26, 0); -lean_inc(x_75); -lean_dec(x_26); -if (lean_is_scalar(x_74)) { - x_76 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_99)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_76 = x_74; + x_100 = x_99; } -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_73); -return x_76; +lean_ctor_set(x_100, 0, x_43); +lean_ctor_set(x_100, 1, x_98); +return x_100; } } +block_168: +{ +if (x_112 == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; uint8_t x_116; +x_113 = lean_st_ref_take(x_103, x_109); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec(x_113); +x_116 = !lean_is_exclusive(x_114); +if (x_116 == 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; +x_117 = lean_ctor_get(x_114, 0); +x_118 = lean_ctor_get(x_114, 4); +lean_dec(x_118); +x_119 = l_Lean_Kernel_enableDiag(x_117, x_106); +lean_ctor_set(x_114, 4, x_39); +lean_ctor_set(x_114, 0, x_119); +x_120 = lean_st_ref_set(x_103, x_114, x_115); +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +lean_dec(x_120); +x_122 = lean_box(0); +lean_inc(x_103); +x_123 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_28, x_106, x_1, x_122, x_33, x_103, x_121); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_st_ref_get(x_103, x_125); +lean_dec(x_103); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_129, 0, x_124); +lean_ctor_set(x_129, 1, x_127); +x_41 = x_129; +x_42 = x_128; +goto block_101; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_77 = lean_ctor_get(x_29, 1); -x_78 = lean_ctor_get(x_29, 2); -x_79 = lean_ctor_get(x_29, 3); -x_80 = lean_ctor_get(x_29, 4); -x_81 = lean_ctor_get(x_29, 6); -lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_dec(x_29); -x_82 = l_Lean_PersistentArray_append___rarg(x_77, x_34); -x_83 = lean_ctor_get_uint8(x_81, sizeof(void*)*2); -x_84 = lean_ctor_get(x_81, 0); -lean_inc(x_84); -x_85 = lean_ctor_get(x_81, 1); -lean_inc(x_85); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_86 = x_81; -} else { - lean_dec_ref(x_81); - x_86 = lean_box(0); -} -x_87 = lean_ctor_get(x_35, 1); -lean_inc(x_87); -lean_dec(x_35); -x_88 = l_Lean_PersistentArray_append___rarg(x_85, x_87); -if (lean_is_scalar(x_86)) { - x_89 = lean_alloc_ctor(0, 2, 1); -} else { - x_89 = x_86; +uint8_t x_130; +lean_dec(x_103); +x_130 = !lean_is_exclusive(x_123); +if (x_130 == 0) +{ +return x_123; } -lean_ctor_set(x_89, 0, x_84); -lean_ctor_set(x_89, 1, x_88); -lean_ctor_set_uint8(x_89, sizeof(void*)*2, x_83); -x_90 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_33); -x_91 = lean_alloc_ctor(0, 8, 0); -lean_ctor_set(x_91, 0, x_31); -lean_ctor_set(x_91, 1, x_82); -lean_ctor_set(x_91, 2, x_78); -lean_ctor_set(x_91, 3, x_79); -lean_ctor_set(x_91, 4, x_80); -lean_ctor_set(x_91, 5, x_32); -lean_ctor_set(x_91, 6, x_89); -lean_ctor_set(x_91, 7, x_90); -x_92 = lean_st_ref_set(x_3, x_91, x_30); -if (lean_obj_tag(x_26) == 0) +else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_123, 0); +x_132 = lean_ctor_get(x_123, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_123); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +return x_133; } -x_95 = lean_ctor_get(x_26, 0); -lean_inc(x_95); -lean_dec(x_26); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(1, 2, 0); -} else { - x_96 = x_94; - lean_ctor_set_tag(x_96, 1); } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_97 = lean_ctor_get(x_92, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_98 = x_92; +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_134 = lean_ctor_get(x_114, 0); +x_135 = lean_ctor_get(x_114, 1); +x_136 = lean_ctor_get(x_114, 2); +x_137 = lean_ctor_get(x_114, 3); +x_138 = lean_ctor_get(x_114, 5); +x_139 = lean_ctor_get(x_114, 6); +lean_inc(x_139); +lean_inc(x_138); +lean_inc(x_137); +lean_inc(x_136); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_114); +x_140 = l_Lean_Kernel_enableDiag(x_134, x_106); +x_141 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_135); +lean_ctor_set(x_141, 2, x_136); +lean_ctor_set(x_141, 3, x_137); +lean_ctor_set(x_141, 4, x_39); +lean_ctor_set(x_141, 5, x_138); +lean_ctor_set(x_141, 6, x_139); +x_142 = lean_st_ref_set(x_103, x_141, x_115); +x_143 = lean_ctor_get(x_142, 1); +lean_inc(x_143); +lean_dec(x_142); +x_144 = lean_box(0); +lean_inc(x_103); +x_145 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_28, x_106, x_1, x_144, x_33, x_103, x_143); +if (lean_obj_tag(x_145) == 0) +{ +lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +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_st_ref_get(x_103, x_147); +lean_dec(x_103); +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_146); +lean_ctor_set(x_151, 1, x_149); +x_41 = x_151; +x_42 = x_150; +goto block_101; +} +else +{ +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +lean_dec(x_103); +x_152 = lean_ctor_get(x_145, 0); +lean_inc(x_152); +x_153 = lean_ctor_get(x_145, 1); +lean_inc(x_153); +if (lean_is_exclusive(x_145)) { + lean_ctor_release(x_145, 0); + lean_ctor_release(x_145, 1); + x_154 = x_145; } else { - lean_dec_ref(x_92); - x_98 = lean_box(0); + lean_dec_ref(x_145); + x_154 = lean_box(0); } -x_99 = lean_ctor_get(x_26, 0); -lean_inc(x_99); -lean_dec(x_26); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_154)) { + x_155 = lean_alloc_ctor(1, 2, 0); } else { - x_100 = x_98; + x_155 = x_154; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_155, 0, x_152); +lean_ctor_set(x_155, 1, x_153); +return x_155; } } } +else +{ +lean_object* x_156; lean_object* x_157; +x_156 = lean_box(0); +lean_inc(x_103); +x_157 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_28, x_106, x_1, x_156, x_33, x_103, x_109); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; +x_158 = lean_ctor_get(x_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_st_ref_get(x_103, x_159); +lean_dec(x_103); +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_158); +lean_ctor_set(x_163, 1, x_161); +x_41 = x_163; +x_42 = x_162; +goto block_101; } else { -uint8_t x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_131 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); -lean_dec(x_12); -x_132 = l_Lean_Elab_Command_State_infoState___default___closed__3; -x_133 = l_Lean_Elab_Command_State_messages___default___closed__3; -x_134 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -lean_ctor_set_uint8(x_134, sizeof(void*)*2, x_131); -x_135 = l_Lean_Elab_Command_State_nextMacroScope___default___closed__1; -x_136 = l_Lean_Elab_Command_liftCoreM___rarg___closed__1; -x_137 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_137, 0, x_13); -lean_ctor_set(x_137, 1, x_135); -lean_ctor_set(x_137, 2, x_14); -lean_ctor_set(x_137, 3, x_15); -lean_ctor_set(x_137, 4, x_136); -lean_ctor_set(x_137, 5, x_133); -lean_ctor_set(x_137, 6, x_134); -x_176 = lean_st_mk_ref(x_137, x_10); -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); -lean_inc(x_178); -lean_dec(x_176); -lean_inc(x_177); -x_179 = lean_apply_3(x_1, x_11, x_177, x_178); -if (lean_obj_tag(x_179) == 0) +uint8_t x_164; +lean_dec(x_103); +x_164 = !lean_is_exclusive(x_157); +if (x_164 == 0) { -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); -lean_dec(x_179); -x_182 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_182, 0, x_180); -x_183 = lean_st_ref_get(x_177, x_181); -lean_dec(x_177); -x_184 = lean_ctor_get(x_183, 0); -lean_inc(x_184); -x_185 = lean_ctor_get(x_183, 1); -lean_inc(x_185); -lean_dec(x_183); -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_182); -lean_ctor_set(x_186, 1, x_184); -x_138 = x_186; -x_139 = x_185; -goto block_175; +return x_157; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; uint8_t x_190; -x_187 = lean_ctor_get(x_179, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_179, 1); -lean_inc(x_188); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_189 = x_179; -} else { - lean_dec_ref(x_179); - x_189 = lean_box(0); +lean_object* x_165; lean_object* x_166; lean_object* x_167; +x_165 = lean_ctor_get(x_157, 0); +x_166 = lean_ctor_get(x_157, 1); +lean_inc(x_166); +lean_inc(x_165); +lean_dec(x_157); +x_167 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_167, 0, x_165); +lean_ctor_set(x_167, 1, x_166); +return x_167; +} +} } -x_190 = l_Lean_Exception_isRuntime(x_187); -if (x_190 == 0) +} +} +else { -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_189); -x_191 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_191, 0, x_187); -x_192 = lean_st_ref_get(x_177, x_188); -lean_dec(x_177); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); -lean_inc(x_194); -lean_dec(x_192); -x_195 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_195, 0, x_191); -lean_ctor_set(x_195, 1, x_193); -x_138 = x_195; -x_139 = x_194; -goto block_175; +uint8_t 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_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; uint8_t x_219; uint8_t x_220; +x_171 = lean_ctor_get_uint8(x_16, sizeof(void*)*2); +lean_dec(x_16); +x_172 = l_Lean_Elab_Command_State_infoState___default___closed__3; +x_173 = l_Lean_Elab_Command_State_messages___default___closed__3; +x_174 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_174, 0, x_172); +lean_ctor_set(x_174, 1, x_173); +lean_ctor_set_uint8(x_174, sizeof(void*)*2, x_171); +x_175 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__3; +x_176 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_176, 0, x_20); +lean_ctor_set(x_176, 1, x_13); +lean_ctor_set(x_176, 2, x_15); +lean_ctor_set(x_176, 3, x_17); +lean_ctor_set(x_176, 4, x_175); +lean_ctor_set(x_176, 5, x_173); +lean_ctor_set(x_176, 6, x_174); +x_210 = lean_st_mk_ref(x_176, x_10); +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_210, 1); +lean_inc(x_212); +lean_dec(x_210); +x_213 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__4; +x_214 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_213); +x_215 = lean_st_ref_get(x_211, x_212); +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_ctor_get(x_216, 0); +lean_inc(x_218); +lean_dec(x_216); +x_219 = l_Lean_Kernel_isDiagnosticsEnabled(x_218); +lean_dec(x_218); +if (x_219 == 0) +{ +if (x_214 == 0) +{ +uint8_t x_260; +x_260 = 1; +x_220 = x_260; +goto block_259; } else { -lean_object* x_196; -lean_dec(x_177); -if (lean_is_scalar(x_189)) { - x_196 = lean_alloc_ctor(1, 2, 0); -} else { - x_196 = x_189; +x_220 = x_32; +goto block_259; } -lean_ctor_set(x_196, 0, x_187); -lean_ctor_set(x_196, 1, x_188); -return x_196; } +else +{ +if (x_214 == 0) +{ +x_220 = x_32; +goto block_259; } -block_175: +else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_140 = lean_ctor_get(x_138, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_138, 1); -lean_inc(x_141); -lean_dec(x_138); -x_142 = lean_st_ref_take(x_3, x_139); -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); -lean_inc(x_144); -lean_dec(x_142); -x_145 = lean_ctor_get(x_141, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_141, 2); -lean_inc(x_146); -x_147 = lean_ctor_get(x_141, 3); -lean_inc(x_147); -x_148 = lean_ctor_get(x_141, 5); -lean_inc(x_148); -x_149 = lean_ctor_get(x_141, 6); -lean_inc(x_149); -lean_dec(x_141); -x_150 = lean_ctor_get(x_143, 1); -lean_inc(x_150); -x_151 = lean_ctor_get(x_143, 2); -lean_inc(x_151); -x_152 = lean_ctor_get(x_143, 3); -lean_inc(x_152); -x_153 = lean_ctor_get(x_143, 4); -lean_inc(x_153); -x_154 = lean_ctor_get(x_143, 6); -lean_inc(x_154); -if (lean_is_exclusive(x_143)) { - lean_ctor_release(x_143, 0); - lean_ctor_release(x_143, 1); - lean_ctor_release(x_143, 2); - lean_ctor_release(x_143, 3); - lean_ctor_release(x_143, 4); - lean_ctor_release(x_143, 5); - lean_ctor_release(x_143, 6); - lean_ctor_release(x_143, 7); - x_155 = x_143; +uint8_t x_261; +x_261 = 1; +x_220 = x_261; +goto block_259; +} +} +block_209: +{ +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; uint8_t x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_179 = lean_ctor_get(x_177, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_177, 1); +lean_inc(x_180); +lean_dec(x_177); +x_181 = lean_st_ref_take(x_3, x_178); +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); +x_184 = lean_ctor_get(x_180, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_180, 1); +lean_inc(x_185); +x_186 = lean_ctor_get(x_180, 2); +lean_inc(x_186); +x_187 = lean_ctor_get(x_180, 3); +lean_inc(x_187); +x_188 = lean_ctor_get(x_180, 5); +lean_inc(x_188); +x_189 = lean_ctor_get(x_180, 6); +lean_inc(x_189); +lean_dec(x_180); +x_190 = lean_ctor_get(x_182, 1); +lean_inc(x_190); +x_191 = lean_ctor_get(x_182, 2); +lean_inc(x_191); +x_192 = lean_ctor_get(x_182, 4); +lean_inc(x_192); +x_193 = lean_ctor_get(x_182, 6); +lean_inc(x_193); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + lean_ctor_release(x_182, 2); + lean_ctor_release(x_182, 3); + lean_ctor_release(x_182, 4); + lean_ctor_release(x_182, 5); + lean_ctor_release(x_182, 6); + lean_ctor_release(x_182, 7); + x_194 = x_182; +} else { + lean_dec_ref(x_182); + x_194 = lean_box(0); +} +x_195 = l_Lean_PersistentArray_append___rarg(x_190, x_188); +x_196 = lean_ctor_get_uint8(x_193, sizeof(void*)*2); +x_197 = lean_ctor_get(x_193, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_193, 1); +lean_inc(x_198); +if (lean_is_exclusive(x_193)) { + lean_ctor_release(x_193, 0); + lean_ctor_release(x_193, 1); + x_199 = x_193; } else { - lean_dec_ref(x_143); - x_155 = lean_box(0); + lean_dec_ref(x_193); + x_199 = lean_box(0); } -x_156 = l_Lean_PersistentArray_append___rarg(x_150, x_148); -x_157 = lean_ctor_get_uint8(x_154, sizeof(void*)*2); -x_158 = lean_ctor_get(x_154, 0); -lean_inc(x_158); -x_159 = lean_ctor_get(x_154, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_160 = x_154; +x_200 = lean_ctor_get(x_189, 1); +lean_inc(x_200); +lean_dec(x_189); +x_201 = l_Lean_PersistentArray_append___rarg(x_198, x_200); +if (lean_is_scalar(x_199)) { + x_202 = lean_alloc_ctor(0, 2, 1); +} else { + x_202 = x_199; +} +lean_ctor_set(x_202, 0, x_197); +lean_ctor_set(x_202, 1, x_201); +lean_ctor_set_uint8(x_202, sizeof(void*)*2, x_196); +x_203 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_27, x_187); +if (lean_is_scalar(x_194)) { + x_204 = lean_alloc_ctor(0, 8, 0); +} else { + x_204 = x_194; +} +lean_ctor_set(x_204, 0, x_184); +lean_ctor_set(x_204, 1, x_195); +lean_ctor_set(x_204, 2, x_191); +lean_ctor_set(x_204, 3, x_185); +lean_ctor_set(x_204, 4, x_192); +lean_ctor_set(x_204, 5, x_186); +lean_ctor_set(x_204, 6, x_202); +lean_ctor_set(x_204, 7, x_203); +x_205 = lean_st_ref_set(x_3, x_204, x_183); +x_206 = lean_ctor_get(x_205, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_205)) { + lean_ctor_release(x_205, 0); + lean_ctor_release(x_205, 1); + x_207 = x_205; } else { - lean_dec_ref(x_154); - x_160 = lean_box(0); + lean_dec_ref(x_205); + x_207 = lean_box(0); } -x_161 = lean_ctor_get(x_149, 1); -lean_inc(x_161); -lean_dec(x_149); -x_162 = l_Lean_PersistentArray_append___rarg(x_159, x_161); -if (lean_is_scalar(x_160)) { - x_163 = lean_alloc_ctor(0, 2, 1); +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(0, 2, 0); } else { - x_163 = x_160; + x_208 = x_207; } -lean_ctor_set(x_163, 0, x_158); -lean_ctor_set(x_163, 1, x_162); -lean_ctor_set_uint8(x_163, sizeof(void*)*2, x_157); -x_164 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_2, x_147); -if (lean_is_scalar(x_155)) { - x_165 = lean_alloc_ctor(0, 8, 0); -} else { - x_165 = x_155; -} -lean_ctor_set(x_165, 0, x_145); -lean_ctor_set(x_165, 1, x_156); -lean_ctor_set(x_165, 2, x_151); -lean_ctor_set(x_165, 3, x_152); -lean_ctor_set(x_165, 4, x_153); -lean_ctor_set(x_165, 5, x_146); -lean_ctor_set(x_165, 6, x_163); -lean_ctor_set(x_165, 7, x_164); -x_166 = lean_st_ref_set(x_3, x_165, x_144); -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_167 = lean_ctor_get(x_166, 1); -lean_inc(x_167); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_168 = x_166; +lean_ctor_set(x_208, 0, x_179); +lean_ctor_set(x_208, 1, x_206); +return x_208; +} +block_259: +{ +if (x_220 == 0) +{ +lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; +x_221 = lean_st_ref_take(x_211, x_217); +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +lean_dec(x_221); +x_224 = lean_ctor_get(x_222, 0); +lean_inc(x_224); +x_225 = lean_ctor_get(x_222, 1); +lean_inc(x_225); +x_226 = lean_ctor_get(x_222, 2); +lean_inc(x_226); +x_227 = lean_ctor_get(x_222, 3); +lean_inc(x_227); +x_228 = lean_ctor_get(x_222, 5); +lean_inc(x_228); +x_229 = lean_ctor_get(x_222, 6); +lean_inc(x_229); +if (lean_is_exclusive(x_222)) { + lean_ctor_release(x_222, 0); + lean_ctor_release(x_222, 1); + lean_ctor_release(x_222, 2); + lean_ctor_release(x_222, 3); + lean_ctor_release(x_222, 4); + lean_ctor_release(x_222, 5); + lean_ctor_release(x_222, 6); + x_230 = x_222; +} else { + lean_dec_ref(x_222); + x_230 = lean_box(0); +} +x_231 = l_Lean_Kernel_enableDiag(x_224, x_214); +if (lean_is_scalar(x_230)) { + x_232 = lean_alloc_ctor(0, 7, 0); +} else { + x_232 = x_230; +} +lean_ctor_set(x_232, 0, x_231); +lean_ctor_set(x_232, 1, x_225); +lean_ctor_set(x_232, 2, x_226); +lean_ctor_set(x_232, 3, x_227); +lean_ctor_set(x_232, 4, x_175); +lean_ctor_set(x_232, 5, x_228); +lean_ctor_set(x_232, 6, x_229); +x_233 = lean_st_ref_set(x_211, x_232, x_223); +x_234 = lean_ctor_get(x_233, 1); +lean_inc(x_234); +lean_dec(x_233); +x_235 = lean_box(0); +lean_inc(x_211); +x_236 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_28, x_214, x_1, x_235, x_33, x_211, x_234); +if (lean_obj_tag(x_236) == 0) +{ +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; +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); +x_239 = lean_st_ref_get(x_211, x_238); +lean_dec(x_211); +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 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_242, 0, x_237); +lean_ctor_set(x_242, 1, x_240); +x_177 = x_242; +x_178 = x_241; +goto block_209; +} +else +{ +lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_dec(x_211); +x_243 = lean_ctor_get(x_236, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_236, 1); +lean_inc(x_244); +if (lean_is_exclusive(x_236)) { + lean_ctor_release(x_236, 0); + lean_ctor_release(x_236, 1); + x_245 = x_236; +} else { + lean_dec_ref(x_236); + x_245 = lean_box(0); +} +if (lean_is_scalar(x_245)) { + x_246 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_166); - x_168 = lean_box(0); + x_246 = x_245; +} +lean_ctor_set(x_246, 0, x_243); +lean_ctor_set(x_246, 1, x_244); +return x_246; +} +} +else +{ +lean_object* x_247; lean_object* x_248; +x_247 = lean_box(0); +lean_inc(x_211); +x_248 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_28, x_214, x_1, x_247, x_33, x_211, x_217); +if (lean_obj_tag(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_248, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_248, 1); +lean_inc(x_250); +lean_dec(x_248); +x_251 = lean_st_ref_get(x_211, x_250); +lean_dec(x_211); +x_252 = lean_ctor_get(x_251, 0); +lean_inc(x_252); +x_253 = lean_ctor_get(x_251, 1); +lean_inc(x_253); +lean_dec(x_251); +x_254 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_254, 0, x_249); +lean_ctor_set(x_254, 1, x_252); +x_177 = x_254; +x_178 = x_253; +goto block_209; +} +else +{ +lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; +lean_dec(x_211); +x_255 = lean_ctor_get(x_248, 0); +lean_inc(x_255); +x_256 = lean_ctor_get(x_248, 1); +lean_inc(x_256); +if (lean_is_exclusive(x_248)) { + lean_ctor_release(x_248, 0); + lean_ctor_release(x_248, 1); + x_257 = x_248; +} else { + lean_dec_ref(x_248); + x_257 = lean_box(0); +} +if (lean_is_scalar(x_257)) { + x_258 = lean_alloc_ctor(1, 2, 0); +} else { + x_258 = x_257; +} +lean_ctor_set(x_258, 0, x_255); +lean_ctor_set(x_258, 1, x_256); +return x_258; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__5(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +return x_9; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___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_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +lean_inc(x_2); +x_5 = lean_apply_3(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +uint8_t x_6; +lean_dec(x_2); +x_6 = !lean_is_exclusive(x_5); +if (x_6 == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_5, 0); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_5, 0, x_8); +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +} +else +{ +uint8_t x_13; +x_13 = !lean_is_exclusive(x_5); +if (x_13 == 0) +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_ctor_get(x_5, 0); +x_15 = l_Lean_Exception_isRuntime(x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_2); +x_16 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_16, 0, x_14); +lean_ctor_set_tag(x_5, 0); +lean_ctor_set(x_5, 0, x_16); +return x_5; +} +else +{ +uint8_t x_17; +x_17 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); +lean_dec(x_2); +if (x_17 == 0) +{ +return x_5; +} +else +{ +lean_object* x_18; +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_14); +lean_ctor_set_tag(x_5, 0); +lean_ctor_set(x_5, 0, x_18); +return x_5; } -x_169 = lean_ctor_get(x_140, 0); -lean_inc(x_169); -lean_dec(x_140); -if (lean_is_scalar(x_168)) { - x_170 = lean_alloc_ctor(1, 2, 0); -} else { - x_170 = x_168; - lean_ctor_set_tag(x_170, 1); } -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_167); -return x_170; } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_171 = lean_ctor_get(x_166, 1); -lean_inc(x_171); -if (lean_is_exclusive(x_166)) { - lean_ctor_release(x_166, 0); - lean_ctor_release(x_166, 1); - x_172 = x_166; -} else { - lean_dec_ref(x_166); - x_172 = lean_box(0); -} -x_173 = lean_ctor_get(x_140, 0); -lean_inc(x_173); -lean_dec(x_140); -if (lean_is_scalar(x_172)) { - x_174 = lean_alloc_ctor(0, 2, 0); -} else { - x_174 = x_172; +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_5); +x_21 = l_Lean_Exception_isRuntime(x_19); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_2); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +return x_23; } -lean_ctor_set(x_174, 0, x_173); -lean_ctor_set(x_174, 1, x_171); -return x_174; +else +{ +uint8_t x_24; +x_24 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); +lean_dec(x_2); +if (x_24 == 0) +{ +lean_object* x_25; +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_19); +lean_ctor_set(x_25, 1, x_20); +return x_25; } +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_26, 0, x_19); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_20); +return x_27; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftCoreM___rarg___boxed), 4, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__3(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftCoreM___rarg___lambda__1), 4, 1); +lean_closure_set(x_5, 0, x_1); +x_6 = l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg(x_5, x_2, x_3, x_4); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg(x_7, x_2, x_3, x_8); +lean_dec(x_7); +return x_9; } +else +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_6); +if (x_10 == 0) +{ +return x_6; } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__4(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_6, 0); +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_6); +x_13 = lean_alloc_ctor(1, 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_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Command_liftCoreM___spec__2(x_1, x_2); -lean_dec(x_1); -return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM(lean_object* x_1) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftCoreM___spec__5(x_1, x_5, x_6, x_4); -lean_dec(x_1); -return x_7; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftCoreM___rarg___boxed), 4, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_3; -x_3 = l_Lean_PersistentArray_mapM___at_Lean_Elab_Command_liftCoreM___spec__1(x_1, x_2); +lean_object* x_5; +x_5 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_liftCoreM___spec__1___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); lean_dec(x_1); -return x_3; +return x_5; } } LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { @@ -6560,32 +7004,6 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_apply_1(x_1, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Command_liftEIO___rarg___boxed), 4, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftEIO___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_Elab_Command_liftEIO___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftIO___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -8836,119 +9254,7 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_2; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_5 = lean_io_mono_nanos_now(x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_apply_3(x_1, x_2, x_3, x_7); -if (lean_obj_tag(x_8) == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_8, 1); -lean_inc(x_10); -lean_dec(x_8); -x_11 = lean_io_mono_nanos_now(x_10); -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; lean_object* x_15; double x_16; double x_17; double x_18; double x_19; double x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_13 = lean_ctor_get(x_11, 0); -x_14 = 0; -x_15 = lean_unsigned_to_nat(0u); -x_16 = l_Float_ofScientific(x_6, x_14, x_15); -lean_dec(x_6); -x_17 = l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; -x_18 = lean_float_div(x_16, x_17); -x_19 = l_Float_ofScientific(x_13, x_14, x_15); -lean_dec(x_13); -x_20 = lean_float_div(x_19, x_17); -x_21 = lean_box_float(x_18); -x_22 = lean_box_float(x_20); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_9); -lean_ctor_set(x_24, 1, x_23); -lean_ctor_set(x_11, 0, x_24); -return x_11; -} -else -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; double x_29; double x_30; double x_31; double x_32; double x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_11); -x_27 = 0; -x_28 = lean_unsigned_to_nat(0u); -x_29 = l_Float_ofScientific(x_6, x_27, x_28); -lean_dec(x_6); -x_30 = l_Lean_withStartStopSeconds___at_Lean_Elab_Command_runLinters___spec__7___closed__1; -x_31 = lean_float_div(x_29, x_30); -x_32 = l_Float_ofScientific(x_25, x_27, x_28); -lean_dec(x_25); -x_33 = lean_float_div(x_32, x_30); -x_34 = lean_box_float(x_31); -x_35 = lean_box_float(x_33); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_9); -lean_ctor_set(x_37, 1, x_36); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_26); -return x_38; -} -} -else -{ -uint8_t x_39; -lean_dec(x_6); -x_39 = !lean_is_exclusive(x_8); -if (x_39 == 0) -{ -return x_8; -} -else -{ -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_8, 0); -x_41 = lean_ctor_get(x_8, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_8); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; -} -} -} -} -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, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___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) { _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -9206,124 +9512,58 @@ if (lean_is_exclusive(x_92)) { x_94 = lean_box(0); } x_95 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_94; -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; -} -} -} -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) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_4); -return x_6; -} -else -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_4); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_apply_3(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_5, 0, x_8); -return x_5; +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(0, 2, 0); +} else { + x_96 = x_94; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_5); -x_11 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_11, 0, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } } -else +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) +if (lean_obj_tag(x_1) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set_tag(x_5, 0); -lean_ctor_set(x_5, 0, x_15); -return x_5; +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_5, 0); -x_17 = lean_ctor_get(x_5, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_5); -x_18 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_18, 0, x_16); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -return x_19; -} +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__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_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_inc(x_7); -x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__8(x_1, x_5, x_2, x_3, x_7, x_8, x_9); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7(x_1, x_5, x_2, x_3, x_7, x_8, x_9); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__9(x_4, x_7, x_8, x_11); +x_12 = l_MonadExcept_ofExcept___at_Lean_Elab_Command_runLinters___spec__8(x_4, x_7, x_8, x_11); lean_dec(x_7); lean_dec(x_4); return x_12; } } -static double _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -9333,7 +9573,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -9341,11 +9581,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { double x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1; +x_15 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_16 = lean_alloc_ctor(0, 2, 17); @@ -9354,7 +9594,7 @@ lean_ctor_set(x_16, 1, x_3); lean_ctor_set_float(x_16, sizeof(void*)*2, x_15); lean_ctor_set_float(x_16, sizeof(void*)*2 + 8, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 16, x_2); -x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__2; +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__2; x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_17); lean_dec(x_7); if (x_18 == 0) @@ -9365,7 +9605,7 @@ lean_object* x_19; lean_object* x_20; lean_dec(x_3); lean_dec(x_1); x_19 = lean_box(0); -x_20 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); +x_20 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); return x_20; } else @@ -9379,7 +9619,7 @@ lean_ctor_set_float(x_21, sizeof(void*)*2, x_9); lean_ctor_set_float(x_21, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 16, x_2); x_22 = lean_box(0); -x_23 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); +x_23 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); return x_23; } } @@ -9394,12 +9634,12 @@ lean_ctor_set_float(x_24, sizeof(void*)*2, x_9); lean_ctor_set_float(x_24, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_24, sizeof(void*)*2 + 16, x_2); x_25 = lean_box(0); -x_26 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); +x_26 = l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__1(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); return x_26; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -9407,16 +9647,16 @@ x_1 = lean_mk_string_from_bytes("", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13() { +static uint8_t _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("internal exception #", 20); -return x_1; +lean_object* x_1; lean_object* x_2; uint8_t x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2___closed__1; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); +return x_3; } } LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_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; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_4, 0); -x_6 = lean_ctor_get(x_4, 3); -x_7 = lean_ctor_get(x_4, 4); -x_8 = lean_ctor_get(x_4, 5); -x_9 = lean_ctor_get(x_4, 6); -x_10 = l_Lean_Core_getMaxHeartbeats(x_6); -x_11 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__1; -x_12 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_11); -x_13 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__2; -x_14 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; -x_15 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__9; -x_16 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__11; +lean_object* x_4; lean_object* x_17; lean_object* x_18; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; uint8_t x_71; +x_39 = lean_box(0); +x_40 = lean_ctor_get(x_1, 0); +lean_inc(x_40); +lean_dec(x_1); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 3); +lean_inc(x_42); +x_43 = lean_ctor_get(x_40, 4); +lean_inc(x_43); +x_44 = lean_ctor_get(x_40, 5); +lean_inc(x_44); +x_45 = lean_ctor_get(x_40, 6); +lean_inc(x_45); +lean_dec(x_40); +x_46 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__3; +x_47 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; +x_48 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__10; +x_49 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__12; +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_46); +lean_ctor_set(x_50, 2, x_45); +lean_ctor_set(x_50, 3, x_47); +lean_ctor_set(x_50, 4, x_48); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_49); +x_51 = lean_io_get_num_heartbeats(x_3); +x_52 = lean_ctor_get(x_51, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_54 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13; +x_55 = l_Lean_Elab_CommandContextInfo_saveNoFileMap___rarg___lambda__1___closed__3; +x_56 = lean_unsigned_to_nat(0u); +x_57 = lean_unsigned_to_nat(1000u); +x_58 = lean_box(0); +x_59 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__2; +x_60 = l_Lean_firstFrontendMacroScope; +x_61 = 0; +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_54); +lean_ctor_set(x_62, 1, x_55); +lean_ctor_set(x_62, 2, x_39); +lean_ctor_set(x_62, 3, x_56); +lean_ctor_set(x_62, 4, x_57); +lean_ctor_set(x_62, 5, x_58); +lean_ctor_set(x_62, 6, x_43); +lean_ctor_set(x_62, 7, x_44); +lean_ctor_set(x_62, 8, x_52); +lean_ctor_set(x_62, 9, x_59); +lean_ctor_set(x_62, 10, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_61); +x_63 = lean_st_mk_ref(x_50, x_53); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_st_ref_get(x_64, x_65); +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, 0); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_Kernel_isDiagnosticsEnabled(x_69); +lean_dec(x_69); +if (x_70 == 0) +{ +uint8_t x_135; +x_135 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14; +if (x_135 == 0) +{ +uint8_t x_136; +x_136 = 1; +x_71 = x_136; +goto block_134; +} +else +{ +x_71 = x_61; +goto block_134; +} +} +else +{ +uint8_t x_137; +x_137 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14; +if (x_137 == 0) +{ +x_71 = x_61; +goto block_134; +} +else +{ +uint8_t x_138; +x_138 = 1; +x_71 = x_138; +goto block_134; +} +} +block_16: +{ +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +lean_ctor_set(x_4, 0, x_7); +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); lean_inc(x_9); -lean_inc(x_5); -x_17 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_17, 0, x_5); -lean_ctor_set(x_17, 1, x_13); -lean_ctor_set(x_17, 2, x_9); -lean_ctor_set(x_17, 3, x_14); -lean_ctor_set(x_17, 4, x_15); -lean_ctor_set(x_17, 5, x_14); -lean_ctor_set(x_17, 6, x_16); -x_18 = lean_io_get_num_heartbeats(x_3); -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_ContextInfo_runCoreM___rarg___closed__12; -x_22 = l_Lean_Elab_CommandContextInfo_saveNoFileMap___rarg___lambda__1___closed__3; -x_23 = lean_unsigned_to_nat(0u); -x_24 = lean_unsigned_to_nat(1000u); -x_25 = lean_box(0); -x_26 = l_Lean_firstFrontendMacroScope; -x_27 = 0; lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_28 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_28, 0, x_21); -lean_ctor_set(x_28, 1, x_22); -lean_ctor_set(x_28, 2, x_6); -lean_ctor_set(x_28, 3, x_23); -lean_ctor_set(x_28, 4, x_24); -lean_ctor_set(x_28, 5, x_25); -lean_ctor_set(x_28, 6, x_7); -lean_ctor_set(x_28, 7, x_8); -lean_ctor_set(x_28, 8, x_19); -lean_ctor_set(x_28, 9, x_10); -lean_ctor_set(x_28, 10, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*11, x_27); -lean_ctor_set_uint8(x_28, sizeof(void*)*11 + 1, x_12); -x_29 = lean_st_mk_ref(x_17, x_20); -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_30); -x_32 = lean_apply_3(x_2, x_28, x_30, x_31); -if (lean_obj_tag(x_32) == 0) +lean_dec(x_4); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_st_ref_get(x_30, x_34); -lean_dec(x_30); -x_36 = !lean_is_exclusive(x_35); -if (x_36 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) { -lean_object* x_37; -x_37 = lean_ctor_get(x_35, 0); -lean_dec(x_37); -lean_ctor_set(x_35, 0, x_33); -return x_35; +return x_4; } else { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_35, 1); -lean_inc(x_38); -lean_dec(x_35); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_33); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} } } +block_38: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_MessageData_toString(x_19, x_18); +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_alloc_ctor(18, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set_tag(x_20, 1); +lean_ctor_set(x_20, 0, x_23); +x_4 = x_20; +goto block_16; +} else { -lean_object* x_40; -lean_dec(x_30); -x_40 = lean_ctor_get(x_32, 0); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) +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_alloc_ctor(18, 1, 0); +lean_ctor_set(x_26, 0, x_24); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_4 = x_27; +goto block_16; +} +} +else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_32, 1); -lean_inc(x_41); -lean_dec(x_32); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_MessageData_toString(x_42, x_41); -if (lean_obj_tag(x_43) == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) { -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +x_4 = x_20; +goto block_16; +} +else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set_tag(x_43, 1); -lean_ctor_set(x_43, 0, x_46); -return x_43; +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); +x_4 = x_31; +goto block_16; +} +} } 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 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_49, 0, x_47); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_48); -return x_50; +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_32 = lean_ctor_get(x_17, 0); +lean_inc(x_32); +lean_dec(x_17); +x_33 = l___private_Init_Data_Repr_0__Nat_reprFast(x_32); +x_34 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__1; +x_35 = lean_string_append(x_34, x_33); +lean_dec(x_33); +x_36 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_18); +x_4 = x_37; +goto block_16; +} +} +block_134: +{ +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_st_ref_take(x_64, x_68); +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_is_exclusive(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_76 = lean_ctor_get(x_73, 0); +x_77 = lean_ctor_get(x_73, 4); +lean_dec(x_77); +x_78 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14; +x_79 = l_Lean_Kernel_enableDiag(x_76, x_78); +lean_ctor_set(x_73, 4, x_48); +lean_ctor_set(x_73, 0, x_79); +x_80 = lean_st_ref_set(x_64, x_73, x_74); +x_81 = lean_ctor_get(x_80, 1); +lean_inc(x_81); +lean_dec(x_80); +x_82 = lean_box(0); +lean_inc(x_64); +x_83 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2(x_39, x_78, x_42, x_2, x_48, x_82, x_62, x_64, x_81); +if (lean_obj_tag(x_83) == 0) +{ +lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_84 = lean_ctor_get(x_83, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_st_ref_get(x_64, x_85); +lean_dec(x_64); +x_87 = !lean_is_exclusive(x_86); +if (x_87 == 0) +{ +lean_object* x_88; lean_object* x_89; +x_88 = lean_ctor_get(x_86, 0); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_84); +lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_86, 0, x_89); +x_4 = x_86; +goto block_16; +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_90 = lean_ctor_get(x_86, 0); +x_91 = lean_ctor_get(x_86, 1); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_86); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_84); +lean_ctor_set(x_92, 1, x_90); +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_91); +x_4 = x_93; +goto block_16; } } else { -uint8_t x_51; -x_51 = !lean_is_exclusive(x_43); -if (x_51 == 0) +lean_object* x_94; lean_object* x_95; +lean_dec(x_64); +x_94 = lean_ctor_get(x_83, 0); +lean_inc(x_94); +x_95 = lean_ctor_get(x_83, 1); +lean_inc(x_95); +lean_dec(x_83); +x_17 = x_94; +x_18 = x_95; +goto block_38; +} +} +else { -return x_43; +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; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_96 = lean_ctor_get(x_73, 0); +x_97 = lean_ctor_get(x_73, 1); +x_98 = lean_ctor_get(x_73, 2); +x_99 = lean_ctor_get(x_73, 3); +x_100 = lean_ctor_get(x_73, 5); +x_101 = lean_ctor_get(x_73, 6); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); +lean_inc(x_96); +lean_dec(x_73); +x_102 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14; +x_103 = l_Lean_Kernel_enableDiag(x_96, x_102); +x_104 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_97); +lean_ctor_set(x_104, 2, x_98); +lean_ctor_set(x_104, 3, x_99); +lean_ctor_set(x_104, 4, x_48); +lean_ctor_set(x_104, 5, x_100); +lean_ctor_set(x_104, 6, x_101); +x_105 = lean_st_ref_set(x_64, x_104, x_74); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); +x_107 = lean_box(0); +lean_inc(x_64); +x_108 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2(x_39, x_102, x_42, x_2, x_48, x_107, x_62, x_64, x_106); +if (lean_obj_tag(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_object* x_116; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_st_ref_get(x_64, x_110); +lean_dec(x_64); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; +} else { + lean_dec_ref(x_111); + x_114 = lean_box(0); +} +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_109); +lean_ctor_set(x_115, 1, x_112); +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); +x_4 = x_116; +goto block_16; } 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_117; lean_object* x_118; +lean_dec(x_64); +x_117 = lean_ctor_get(x_108, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_108, 1); +lean_inc(x_118); +lean_dec(x_108); +x_17 = x_117; +x_18 = x_118; +goto block_38; } } } else { -uint8_t x_55; -x_55 = !lean_is_exclusive(x_32); -if (x_55 == 0) +uint8_t x_119; lean_object* x_120; lean_object* x_121; +x_119 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14; +x_120 = lean_box(0); +lean_inc(x_64); +x_121 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2(x_39, x_119, x_42, x_2, x_48, x_120, x_62, x_64, x_68); +if (lean_obj_tag(x_121) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_56 = lean_ctor_get(x_32, 0); -lean_dec(x_56); -x_57 = lean_ctor_get(x_40, 0); -lean_inc(x_57); -lean_dec(x_40); -x_58 = l___private_Init_Data_Repr_0__Nat_reprFast(x_57); -x_59 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13; -x_60 = lean_string_append(x_59, x_58); -lean_dec(x_58); -x_61 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_32, 0, x_61); -return x_32; +lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; +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_st_ref_get(x_64, x_123); +lean_dec(x_64); +x_125 = !lean_is_exclusive(x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_124, 0); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_122); +lean_ctor_set(x_127, 1, x_126); +lean_ctor_set(x_124, 0, x_127); +x_4 = x_124; +goto block_16; } 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; -x_62 = lean_ctor_get(x_32, 1); -lean_inc(x_62); -lean_dec(x_32); -x_63 = lean_ctor_get(x_40, 0); -lean_inc(x_63); -lean_dec(x_40); -x_64 = l___private_Init_Data_Repr_0__Nat_reprFast(x_63); -x_65 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13; -x_66 = lean_string_append(x_65, x_64); +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_128 = lean_ctor_get(x_124, 0); +x_129 = lean_ctor_get(x_124, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_124); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_122); +lean_ctor_set(x_130, 1, x_128); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_129); +x_4 = x_131; +goto block_16; +} +} +else +{ +lean_object* x_132; lean_object* x_133; lean_dec(x_64); -x_67 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_67, 0, x_66); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_62); -return x_68; +x_132 = lean_ctor_get(x_121, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_121, 1); +lean_inc(x_133); +lean_dec(x_121); +x_17 = x_132; +x_18 = x_133; +goto block_38; } } } @@ -2506,17 +3168,29 @@ LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runCoreM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_ContextInfo_runCoreM___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_ContextInfo_runCoreM___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runCoreM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runCoreM___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) { _start: { -lean_object* x_4; -x_4 = l_Lean_Elab_ContextInfo_runCoreM___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runCoreM___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) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; } } static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__1() { @@ -2548,7 +3222,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__2() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2568,7 +3242,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__4() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2640,7 +3314,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__11() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2674,7 +3348,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__14() _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2718,7 +3392,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__17() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__9; x_2 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -2745,12 +3419,14 @@ lean_ctor_set(x_10, 5, x_5); lean_ctor_set_uint8(x_10, sizeof(void*)*6, x_9); lean_ctor_set_uint8(x_10, sizeof(void*)*6 + 1, x_9); x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); x_12 = lean_ctor_get(x_11, 2); +lean_inc(x_12); +lean_dec(x_11); x_13 = lean_box(0); x_14 = l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__16; -x_15 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_15 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_16 = l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__17; -lean_inc(x_12); x_17 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_17, 0, x_12); lean_ctor_set(x_17, 1, x_14); @@ -2821,19 +3497,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runMetaM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_ContextInfo_runMetaM___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_ContextInfo_runMetaM___rarg), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_runMetaM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_1); -return x_5; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_toPPContext(lean_object* x_1, lean_object* x_2) { _start: { @@ -3304,19 +3971,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_TermInfo_runMetaM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Elab_TermInfo_runMetaM___rarg___boxed), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_TermInfo_runMetaM___rarg), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_TermInfo_runMetaM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_2); -return x_5; -} -} static lean_object* _init_l_Lean_Elab_TermInfo_format___lambda__1___closed__1() { _start: { @@ -3800,7 +4458,6 @@ x_6 = lean_alloc_closure((void*)(l_Lean_Elab_TermInfo_format___lambda__2), 7, 2) lean_closure_set(x_6, 0, x_5); lean_closure_set(x_6, 1, x_4); x_7 = l_Lean_Elab_TermInfo_runMetaM___rarg(x_2, x_1, x_6, x_3); -lean_dec(x_1); return x_7; } } @@ -4153,7 +4810,6 @@ lean_closure_set(x_35, 2, x_32); lean_closure_set(x_35, 3, x_34); lean_closure_set(x_35, 4, x_2); x_36 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_1, x_33, x_35, x_3); -lean_dec(x_1); return x_36; } default: @@ -4572,7 +5228,6 @@ lean_closure_set(x_6, 0, x_5); lean_closure_set(x_6, 1, x_2); lean_closure_set(x_6, 2, x_1); x_7 = l_Lean_Elab_ContextInfo_runMetaM___rarg(x_1, x_4, x_6, x_3); -lean_dec(x_1); return x_7; } } @@ -4863,7 +5518,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_ppGoals___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4876,7 +5531,7 @@ static lean_object* _init_l_Lean_Elab_ContextInfo_ppGoals___closed__2() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_ContextInfo_ppGoals___closed__1; -x_2 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_2 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -4921,6 +5576,7 @@ else { lean_object* x_9; lean_object* x_10; lean_dec(x_2); +lean_dec(x_1); x_9 = l_Lean_Elab_ContextInfo_ppGoals___closed__4; x_10 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_10, 0, x_9); @@ -4938,15 +5594,6 @@ lean_dec(x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_ContextInfo_ppGoals___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Elab_ContextInfo_ppGoals(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} static lean_object* _init_l_Lean_Elab_TacticInfo_format___closed__1() { _start: { @@ -5050,7 +5697,6 @@ lean_ctor_set(x_18, 1, x_5); x_19 = lean_ctor_get(x_2, 2); lean_inc(x_19); x_20 = l_Lean_Elab_ContextInfo_ppGoals(x_15, x_19, x_3); -lean_dec(x_15); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; @@ -5062,7 +5708,6 @@ lean_dec(x_20); x_23 = lean_ctor_get(x_2, 4); lean_inc(x_23); x_24 = l_Lean_Elab_ContextInfo_ppGoals(x_18, x_23, x_22); -lean_dec(x_18); if (lean_obj_tag(x_24) == 0) { uint8_t x_25; @@ -5276,7 +5921,6 @@ lean_ctor_set(x_89, 1, x_5); x_90 = lean_ctor_get(x_2, 2); lean_inc(x_90); x_91 = l_Lean_Elab_ContextInfo_ppGoals(x_86, x_90, x_3); -lean_dec(x_86); if (lean_obj_tag(x_91) == 0) { lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; @@ -5288,7 +5932,6 @@ lean_dec(x_91); x_94 = lean_ctor_get(x_2, 4); lean_inc(x_94); x_95 = l_Lean_Elab_ContextInfo_ppGoals(x_89, x_94, x_93); -lean_dec(x_89); if (lean_obj_tag(x_95) == 0) { lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; @@ -5607,7 +6250,7 @@ static lean_object* _init_l_Lean_Elab_UserWidgetInfo_format___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__7; +x_1 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__8; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6864,7 +7507,7 @@ if (x_2 == 0) lean_object* x_3; lean_object* x_4; x_3 = lean_ctor_get(x_1, 1); lean_dec(x_3); -x_4 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_4 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; lean_ctor_set(x_1, 1, x_4); return x_1; } @@ -6875,7 +7518,7 @@ x_5 = lean_ctor_get_uint8(x_1, sizeof(void*)*2); x_6 = lean_ctor_get(x_1, 0); lean_inc(x_6); lean_dec(x_1); -x_7 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_7 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_8 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_8, 0, x_6); lean_ctor_set(x_8, 1, x_7); @@ -7080,7 +7723,7 @@ return x_9; else { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_10 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_11 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_11, 0, x_2); lean_ctor_set(x_11, 1, x_10); @@ -7495,7 +8138,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_16 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_1); lean_ctor_set(x_17, 1, x_16); @@ -8230,7 +8873,7 @@ lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_5, 1); lean_inc(x_15); lean_dec(x_5); -x_16 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__5; +x_16 = l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__6; x_17 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_17, 0, x_1); lean_ctor_set(x_17, 1, x_16); @@ -10839,7 +11482,7 @@ static lean_object* _init_l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___clos _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("( __do_lift._@.Lean.Elab.InfoTree.Main._hyg.4310.0 ).isNone\n ", 62); +x_1 = lean_mk_string_from_bytes("( __do_lift._@.Lean.Elab.InfoTree.Main._hyg.4318.0 ).isNone\n ", 62); return x_1; } } @@ -10867,7 +11510,7 @@ static lean_object* _init_l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___clos lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_PartialContextInfo_mergeIntoOuter_x3f___closed__1; x_2 = l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___closed__4; -x_3 = lean_unsigned_to_nat(401u); +x_3 = lean_unsigned_to_nat(402u); x_4 = lean_unsigned_to_nat(2u); x_5 = l_Lean_Elab_assignInfoHoleId___rarg___lambda__2___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -11498,6 +12141,10 @@ l_Lean_PersistentArray_findSomeMAux___at_Lean_Elab_InfoTree_findInfo_x3f___spec_ lean_mark_persistent(l_Lean_PersistentArray_findSomeMAux___at_Lean_Elab_InfoTree_findInfo_x3f___spec__2___closed__3); l_Lean_PersistentHashMap_findAux___at_Lean_Elab_InfoTree_substitute___spec__7___closed__1 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_InfoTree_substitute___spec__7___closed__1(); l_Lean_PersistentHashMap_findAux___at_Lean_Elab_InfoTree_substitute___spec__7___closed__2 = _init_l_Lean_PersistentHashMap_findAux___at_Lean_Elab_InfoTree_substitute___spec__7___closed__2(); +l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__1___closed__1 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__1___closed__1); +l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2___closed__1 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_ContextInfo_runCoreM___rarg___lambda__2___closed__1); l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__1 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__1); l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__2 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__2(); @@ -11524,6 +12171,7 @@ l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__12 = _init_l_Lean_Elab_Context lean_mark_persistent(l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__12); l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13(); lean_mark_persistent(l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__13); +l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14 = _init_l_Lean_Elab_ContextInfo_runCoreM___rarg___closed__14(); l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__1 = _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__1); l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__2 = _init_l_Lean_Elab_ContextInfo_runMetaM___rarg___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index b552fcd6bd36..6efad700c53e 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -130,6 +130,8 @@ lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_mkAuxName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___spec__1___rarg___closed__1; +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1; +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18632____closed__7; lean_object* l_Lean_MessageData_joinSep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_reportMatcherResultErrors___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -179,6 +181,7 @@ static lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__2___ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckMatch___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* l_Lean_Meta_getFVarsToGeneralize(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___closed__1; +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___boxed(lean_object**); lean_object* lean_mk_array(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -281,6 +284,7 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMa LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_eraseIndices___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportMatcherResultErrors(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_PrettyPrinter_delabCore___spec__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -293,6 +297,7 @@ lean_object* l_EStateM_instMonad(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitPost___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_mkGeneralizationForbiddenSet_visit___spec__1(lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___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*); @@ -337,10 +342,12 @@ static lean_object* l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambd LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__1___boxed(lean_object*); lean_object* l_Lean_Meta_Match_Pattern_toExpr_visit(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_topSort___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchOptMotive___boxed(lean_object*); extern lean_object* l_instInhabitedPUnit; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabNoMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabNoMatch___closed__1; static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___rarg___closed__1; @@ -394,13 +401,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabInaccessible_declRange___c static lean_object* l_Lean_Meta_transform___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__1___closed__1; lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoMatch___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_main___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getIndexToInclude_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_Term_mkFreshBinderName___at___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchOptMotive(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_toPattern___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch_declRange___closed__3; @@ -419,6 +427,7 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPatte lean_object* l_Lean_mkInaccessible(lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18632____closed__10; +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___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*); @@ -454,14 +463,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoMatch_declRange___closed lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch___closed__2; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__5___rarg___closed__2; lean_object* l_Lean_inaccessible_x3f(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_main___spec__7(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_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(lean_object*, lean_object*, uint8_t, 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_generalize___boxed__const__1; 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*); -uint8_t l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_precheckMatch___lambda__1___closed__1; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__5___lambda__1___boxed(lean_object**); @@ -493,6 +502,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoFun___spec__1___lambda__1___closed__2; lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5; lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withEqs_go___rarg(lean_object*, lean_object*, lean_object*, 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_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___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*); @@ -562,8 +572,10 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Match_0__Lea lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstances___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrs___spec__11___boxed(lean_object**); +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7; 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*); uint8_t l_List_elem___at_Lean_Meta_Occurrences_contains___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__4___closed__2; lean_object* l_Lean_Syntax_mkSep(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabNoFun___closed__3; @@ -610,6 +622,7 @@ lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lea lean_object* l_Lean_Elab_Term_tryPostponeIfMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_waitExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Match_0__Lean_Elab_Term_mkUserNameFor___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__14___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18632____closed__13; @@ -633,6 +646,7 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPat LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__11___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_instantiateAltLHSMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_isMatchUnit_x3f___closed__3; static lean_object* l_Lean_Elab_Term_reportMatcherResultErrors___closed__2; @@ -667,6 +681,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0 lean_object* l_Lean_Meta_isMatchValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_Context_userName___default; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__15(lean_object*, lean_object*, size_t, size_t); +extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at___private_Lean_Meta_WHNF_0__Lean_Meta_cached_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -728,7 +743,6 @@ extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_ToDepElimPattern_main___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___lambda__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13(lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabNoFun___spec__1___lambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabNoFun(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_processVar___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -809,6 +823,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabMatch(lean_object*); lean_object* l_Lean_Expr_replaceFVars(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_ToDepElimPattern_main___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomicDiscrs_x3f_loop___closed__1; +LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___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_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_ToDepElimPattern_normalize_processInaccessible___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_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3; lean_object* l_Lean_Expr_ReplaceImpl_Cache_new(lean_object*); @@ -843,7 +858,6 @@ static lean_object* l_Lean_Elab_Term_ToDepElimPattern_normalize_addVar___closed_ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__6(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__10(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -963,6 +977,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNon LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Term_reportMatcherResultErrors___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchAlts___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandMacrosInPatterns___spec__1(size_t, size_t, lean_object*, 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*); @@ -988,6 +1003,7 @@ static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandNonAtomi size_t lean_usize_sub(size_t, size_t); static lean_object* l_List_forIn_loop___at_Lean_Elab_Term_reportMatcherResultErrors___spec__1___closed__2; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___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*); uint8_t l_Lean_LocalDecl_hasExprMVar(lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_isAtomicDiscr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1171,7 +1187,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabInaccessible(lean_object*, lean_ob static lean_object* l___regBuiltin_Lean_Elab_Term_precheckMatch___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_packMatchTypePatterns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_applyRefMap(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_expandSimpleMatch___closed__10; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Term_precheckMatch___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_Match_0__Lean_Elab_Term_withToClear___spec__7(lean_object*, lean_object*, size_t, size_t); @@ -19736,398 +19752,392 @@ x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_checked___at___private_ return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___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) { _start: { -lean_dec(x_4); +lean_dec(x_5); switch (lean_obj_tag(x_1)) { case 5: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_9 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_7, x_5, x_6); -if (lean_obj_tag(x_9) == 0) +x_10 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_8, x_6, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_8, x_5, x_10); -return x_11; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_9, x_6, x_11); +return x_12; } else { -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_5); +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) { -return x_9; +return x_10; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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_inc(x_13); -lean_dec(x_9); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_dec(x_10); +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; } } } case 6: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 2); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_1, 1); lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 2); +lean_inc(x_18); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_18 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_16, x_5, x_6); -if (lean_obj_tag(x_18) == 0) +x_19 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_17, x_6, x_7); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_17, x_5, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_18, x_6, x_20); +return x_21; } else { -uint8_t x_21; -lean_dec(x_17); -lean_dec(x_5); +uint8_t x_22; +lean_dec(x_18); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) { -return x_18; +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 0); -x_23 = lean_ctor_get(x_18, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_18); -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_19); +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; } } } case 7: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 2); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_1, 1); lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 2); +lean_inc(x_27); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_27 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_25, x_5, x_6); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_26, x_6, x_7); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_26, x_5, x_28); -return x_29; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_27, x_6, x_29); +return x_30; } else { -uint8_t x_30; -lean_dec(x_26); -lean_dec(x_5); +uint8_t x_31; +lean_dec(x_27); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_27); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -return x_27; +return x_28; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 0); -x_32 = lean_ctor_get(x_27, 1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +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_inc(x_31); -lean_dec(x_27); -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_dec(x_28); +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; } } } case 8: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_1, 1); -lean_inc(x_34); -x_35 = lean_ctor_get(x_1, 2); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_1, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_1, 3); +x_36 = lean_ctor_get(x_1, 2); lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 3); +lean_inc(x_37); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_37 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_34, x_5, x_6); -if (lean_obj_tag(x_37) == 0) +x_38 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_35, x_6, x_7); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -lean_inc(x_5); +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +lean_dec(x_38); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_39 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_35, x_5, x_38); -if (lean_obj_tag(x_39) == 0) +x_40 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_36, x_6, x_39); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_36, x_5, x_40); -return x_41; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_37, x_6, x_41); +return x_42; } else { -uint8_t x_42; -lean_dec(x_36); -lean_dec(x_5); +uint8_t x_43; +lean_dec(x_37); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_42 = !lean_is_exclusive(x_39); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_40); +if (x_43 == 0) { -return x_39; +return x_40; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_39, 0); -x_44 = lean_ctor_get(x_39, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 0); +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_39); -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_dec(x_40); +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_46; +uint8_t x_47; +lean_dec(x_37); lean_dec(x_36); -lean_dec(x_35); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_37); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) { -return x_37; +return x_38; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_37, 0); -x_48 = lean_ctor_get(x_37, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_38, 0); +x_49 = lean_ctor_get(x_38, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_37); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_38); +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; } } } case 10: { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_1, 1); -lean_inc(x_50); +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_1, 1); +lean_inc(x_51); lean_dec(x_1); -x_51 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_50, x_5, x_6); -return x_51; +x_52 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_51, x_6, x_7); +return x_52; } case 11: { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_1, 2); -lean_inc(x_52); +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_1, 2); +lean_inc(x_53); lean_dec(x_1); -x_53 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_52, x_5, x_6); -return x_53; +x_54 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_2, x_3, x_4, x_53, x_6, x_7); +return x_54; } default: { -lean_object* x_54; lean_object* x_55; -lean_dec(x_5); +lean_object* x_55; lean_object* x_56; +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_6); -return x_55; +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_7); +return x_56; } } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(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_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___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) { _start: { -lean_object* x_6; -lean_inc(x_3); -x_6 = l_Lean_ForEachExprWhere_visited___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__11___rarg(x_3, x_4, x_5); -if (lean_obj_tag(x_6) == 0) +lean_object* x_7; +lean_inc(x_4); +x_7 = l_Lean_ForEachExprWhere_visited___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__11___rarg(x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_7; uint8_t x_8; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_unbox(x_7); -lean_dec(x_7); -if (x_8 == 0) +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_6, 1); -lean_inc(x_9); -lean_dec(x_6); +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); lean_inc(x_1); -lean_inc(x_3); -x_10 = lean_apply_1(x_1, x_3); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) +lean_inc(x_4); +x_11 = lean_apply_1(x_1, x_4); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_3, x_1, x_2, x_12, x_4, x_9); -return x_13; +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_4, x_1, x_2, x_3, x_13, x_5, x_10); +return x_14; } else { -lean_object* x_14; -lean_inc(x_3); -x_14 = l_Lean_ForEachExprWhere_checked___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__12___rarg(x_3, x_4, x_9); -if (lean_obj_tag(x_14) == 0) +lean_object* x_15; +lean_inc(x_4); +x_15 = l_Lean_ForEachExprWhere_checked___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__12___rarg(x_4, x_5, x_10); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; uint8_t x_16; -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_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_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); lean_inc(x_2); -lean_inc(x_3); -x_18 = lean_apply_2(x_2, x_3, x_17); -if (lean_obj_tag(x_18) == 0) +lean_inc(x_4); +x_19 = lean_apply_2(x_2, x_4, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +if (x_3 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_19, 1); lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_3, x_1, x_2, x_19, x_4, x_20); -return x_21; +lean_dec(x_19); +x_21 = lean_box(0); +x_22 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_4, x_1, x_2, x_3, x_21, x_5, x_20); +return x_22; } else { -uint8_t x_22; +uint8_t x_23; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = !lean_is_exclusive(x_18); -if (x_22 == 0) -{ -return x_18; -} -else +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) { -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; -} -} +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_19, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_19, 0, x_25); +return x_19; } else { lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_14, 1); +x_26 = lean_ctor_get(x_19, 1); lean_inc(x_26); -lean_dec(x_14); +lean_dec(x_19); x_27 = lean_box(0); -x_28 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_3, x_1, x_2, x_27, x_4, 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_26); return x_28; } } +} else { uint8_t x_29; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_14); +x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) { -return x_14; +return x_19; } else { lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_14, 0); -x_31 = lean_ctor_get(x_14, 1); +x_30 = lean_ctor_get(x_19, 0); +x_31 = lean_ctor_get(x_19, 1); lean_inc(x_31); lean_inc(x_30); -lean_dec(x_14); +lean_dec(x_19); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); @@ -20135,62 +20145,100 @@ return x_32; } } } +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_4, x_1, x_2, x_3, x_34, x_5, x_33); +return x_35; +} } else { -uint8_t x_33; +uint8_t x_36; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_6); -if (x_33 == 0) +x_36 = !lean_is_exclusive(x_15); +if (x_36 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_6, 0); -lean_dec(x_34); -x_35 = lean_box(0); -lean_ctor_set(x_6, 0, x_35); -return x_6; +return x_15; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_6, 1); -lean_inc(x_36); -lean_dec(x_6); -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; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_15, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_15); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} } } } else { -uint8_t x_39; +uint8_t x_40; +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_6); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_7); +if (x_40 == 0) { -return x_6; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_7, 0); +lean_dec(x_41); +x_42 = lean_box(0); +lean_ctor_set(x_7, 0, x_42); +return x_7; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_6, 0); -x_41 = lean_ctor_get(x_6, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_6); -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_7, 1); +lean_inc(x_43); +lean_dec(x_7); +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 +{ +uint8_t x_46; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_7); +if (x_46 == 0) +{ +return x_7; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_7, 0); +x_48 = lean_ctor_get(x_7, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_7); +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; } } } @@ -20199,131 +20247,131 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_El _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg), 5, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___boxed), 6, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_ForEachExprWhere_initCache; -x_6 = lean_st_mk_ref(x_5, x_4); -if (lean_obj_tag(x_6) == 0) +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_ForEachExprWhere_initCache; +x_7 = lean_st_mk_ref(x_6, x_5); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -lean_dec(x_6); -lean_inc(x_7); -x_9 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_1, x_2, x_3, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_8); +x_10 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_1, x_2, x_4, x_3, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_st_ref_get(x_7, x_11); -lean_dec(x_7); -if (lean_obj_tag(x_12) == 0) +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_st_ref_get(x_8, x_12); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_10); -return x_12; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -lean_object* x_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_10); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -uint8_t x_17; -lean_dec(x_10); -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) +uint8_t x_18; +lean_dec(x_11); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -return x_12; +return x_13; } 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_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_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_dec(x_13); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } else { -uint8_t x_21; -lean_dec(x_7); -x_21 = !lean_is_exclusive(x_9); -if (x_21 == 0) +uint8_t x_22; +lean_dec(x_8); +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { -return x_9; +return x_10; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_9, 0); -x_23 = lean_ctor_get(x_9, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_9); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_10); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -uint8_t x_25; +uint8_t x_26; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_6); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_7); +if (x_26 == 0) { -return x_6; +return x_7; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_6, 0); -x_27 = lean_ctor_get(x_6, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_7, 0); +x_28 = lean_ctor_get(x_7, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_6); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_7); +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; } } } @@ -20332,7 +20380,7 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_ _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg___boxed), 5, 0); return x_2; } } @@ -20455,7 +20503,7 @@ x_5 = l_Lean_mkHashMapImp___rarg(x_4); x_6 = lean_st_mk_ref(x_5, x_3); if (lean_obj_tag(x_6) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); x_8 = lean_ctor_get(x_6, 1); @@ -20465,62 +20513,63 @@ lean_inc(x_7); x_9 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___lambda__1), 3, 1); lean_closure_set(x_9, 0, x_7); x_10 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___closed__1; -x_11 = l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(x_10, x_9, x_1, x_8); -if (lean_obj_tag(x_11) == 0) +x_11 = 0; +x_12 = l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(x_10, x_9, x_1, x_11, x_8); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = lean_st_ref_get(x_7, x_12); +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = lean_st_ref_get(x_7, x_13); lean_dec(x_7); -return x_13; +return x_14; } else { -uint8_t x_14; +uint8_t x_15; lean_dec(x_7); -x_14 = !lean_is_exclusive(x_11); -if (x_14 == 0) +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) { -return x_11; +return x_12; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_11, 1); +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); +lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_dec(x_12); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; } } } else { -uint8_t x_18; +uint8_t x_19; lean_dec(x_1); -x_18 = !lean_is_exclusive(x_6); -if (x_18 == 0) +x_19 = !lean_is_exclusive(x_6); +if (x_19 == 0) { return x_6; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_6, 0); -x_20 = lean_ctor_get(x_6, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_6, 0); +x_21 = lean_ctor_get(x_6, 1); +lean_inc(x_21); lean_inc(x_20); -lean_inc(x_19); lean_dec(x_6); -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; +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; } } } @@ -20563,6 +20612,36 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1___boxed(lean_object* x_1, 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_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg___lambda__1(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___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) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_ForEachExprWhere_visit_go___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__10___rarg(x_1, x_2, x_7, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg___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_6; lean_object* x_7; +x_6 = lean_unbox(x_4); +lean_dec(x_4); +x_7 = l_Lean_ForEachExprWhere_visit___at___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap_go___spec__9___rarg(x_1, x_2, x_3, x_6, x_5); +return x_7; +} +} LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_ToDepElimPattern_mkPatternRefMap(lean_object* x_1) { _start: { @@ -21717,7 +21796,7 @@ if (x_17 == 0) lean_object* x_18; lean_object* x_19; uint8_t x_20; x_18 = lean_ctor_get(x_16, 0); x_19 = lean_ctor_get(x_16, 1); -x_20 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_18, x_14); +x_20 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_18, x_14); lean_dec(x_18); if (x_20 == 0) { @@ -21771,7 +21850,7 @@ x_34 = lean_ctor_get(x_16, 1); lean_inc(x_34); lean_inc(x_33); lean_dec(x_16); -x_35 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_33, x_14); +x_35 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_33, x_14); lean_dec(x_33); if (x_35 == 0) { @@ -21920,7 +21999,7 @@ lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_17, x_14); +x_19 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_17, x_14); lean_dec(x_17); if (x_19 == 0) { @@ -25416,7 +25495,7 @@ case 1: lean_object* x_11; uint8_t x_12; x_11 = lean_ctor_get(x_2, 0); lean_inc(x_11); -x_12 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_1, x_2); +x_12 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_1, x_2); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; @@ -25781,7 +25860,7 @@ else { uint8_t x_108; lean_dec(x_104); -x_108 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_1, x_2); +x_108 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_1, x_2); if (x_108 == 0) { lean_object* x_109; @@ -26086,7 +26165,7 @@ else { uint8_t x_194; lean_dec(x_190); -x_194 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_1, x_2); +x_194 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_1, x_2); if (x_194 == 0) { lean_object* x_195; lean_object* x_196; @@ -29384,6 +29463,78 @@ goto _start; } } } +static lean_object* _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___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) { +_start: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_13 = lean_ctor_get(x_9, 4); +lean_dec(x_13); +x_14 = lean_ctor_get(x_9, 2); +lean_dec(x_14); +x_15 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1; +x_16 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_15); +lean_ctor_set(x_9, 4, x_16); +lean_ctor_set(x_9, 2, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*11 + 1, x_2); +x_17 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_9, x_10, x_11); +return x_17; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +x_20 = lean_ctor_get(x_9, 3); +x_21 = lean_ctor_get(x_9, 5); +x_22 = lean_ctor_get(x_9, 6); +x_23 = lean_ctor_get(x_9, 7); +x_24 = lean_ctor_get(x_9, 8); +x_25 = lean_ctor_get(x_9, 9); +x_26 = lean_ctor_get(x_9, 10); +x_27 = lean_ctor_get_uint8(x_9, 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_dec(x_9); +x_28 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1; +x_29 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_28); +x_30 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_30, 0, x_18); +lean_ctor_set(x_30, 1, x_19); +lean_ctor_set(x_30, 2, x_1); +lean_ctor_set(x_30, 3, x_20); +lean_ctor_set(x_30, 4, x_29); +lean_ctor_set(x_30, 5, x_21); +lean_ctor_set(x_30, 6, x_22); +lean_ctor_set(x_30, 7, x_23); +lean_ctor_set(x_30, 8, x_24); +lean_ctor_set(x_30, 9, x_25); +lean_ctor_set(x_30, 10, x_26); +lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_27); +lean_ctor_set_uint8(x_30, sizeof(void*)*11 + 1, x_2); +x_31 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_30, x_10, x_11); +return x_31; +} +} +} static lean_object* _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__1() { _start: { @@ -29402,70 +29553,180 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } +static lean_object* _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4() { +_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_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5; +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_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_6); -if (x_9 == 0) +lean_object* x_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_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +x_10 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2; +x_11 = 1; +x_12 = l_Lean_KVMap_setBool(x_9, x_10, x_11); +x_13 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3; +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_13); +x_15 = lean_st_ref_get(x_7, x_8); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Kernel_isDiagnosticsEnabled(x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_10; lean_object* x_11; uint8_t x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_ctor_get(x_6, 2); -x_11 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2; -x_12 = 1; -x_13 = l_Lean_KVMap_setBool(x_10, x_11, x_12); -lean_ctor_set(x_6, 2, x_13); -x_14 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_14; +if (x_14 == 0) +{ +x_20 = x_11; +goto block_48; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t 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; -x_15 = lean_ctor_get(x_6, 0); -x_16 = lean_ctor_get(x_6, 1); -x_17 = lean_ctor_get(x_6, 2); -x_18 = lean_ctor_get(x_6, 3); -x_19 = lean_ctor_get(x_6, 4); -x_20 = lean_ctor_get(x_6, 5); -x_21 = lean_ctor_get(x_6, 6); -x_22 = lean_ctor_get(x_6, 7); -x_23 = lean_ctor_get(x_6, 8); -x_24 = lean_ctor_get(x_6, 9); -x_25 = lean_ctor_get(x_6, 10); -x_26 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); -x_27 = lean_ctor_get_uint8(x_6, sizeof(void*)*11 + 1); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); +uint8_t x_49; +x_49 = 0; +x_20 = x_49; +goto block_48; +} +} +else +{ +if (x_14 == 0) +{ +uint8_t x_50; +x_50 = 0; +x_20 = x_50; +goto block_48; +} +else +{ +x_20 = x_11; +goto block_48; +} +} +block_48: +{ +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_st_ref_take(x_7, x_17); +x_22 = lean_ctor_get(x_21, 0); 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_dec(x_6); -x_28 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2; -x_29 = 1; -x_30 = l_Lean_KVMap_setBool(x_17, x_28, x_29); -x_31 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_31, 0, x_15); -lean_ctor_set(x_31, 1, x_16); -lean_ctor_set(x_31, 2, x_30); -lean_ctor_set(x_31, 3, x_18); -lean_ctor_set(x_31, 4, x_19); -lean_ctor_set(x_31, 5, x_20); -lean_ctor_set(x_31, 6, x_21); -lean_ctor_set(x_31, 7, x_22); -lean_ctor_set(x_31, 8, x_23); -lean_ctor_set(x_31, 9, x_24); -lean_ctor_set(x_31, 10, x_25); -lean_ctor_set_uint8(x_31, sizeof(void*)*11, x_26); -lean_ctor_set_uint8(x_31, sizeof(void*)*11 + 1, x_27); -x_32 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_31, x_7, x_8); +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; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_25 = lean_ctor_get(x_22, 0); +x_26 = lean_ctor_get(x_22, 4); +lean_dec(x_26); +x_27 = l_Lean_Kernel_enableDiag(x_25, x_14); +x_28 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7; +lean_ctor_set(x_22, 4, x_28); +lean_ctor_set(x_22, 0, x_27); +x_29 = lean_st_ref_set(x_7, x_22, x_23); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_box(0); +x_32 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1(x_12, x_14, x_1, x_2, x_3, x_4, x_5, x_31, x_6, x_7, 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; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_33 = lean_ctor_get(x_22, 0); +x_34 = lean_ctor_get(x_22, 1); +x_35 = lean_ctor_get(x_22, 2); +x_36 = lean_ctor_get(x_22, 3); +x_37 = lean_ctor_get(x_22, 5); +x_38 = lean_ctor_get(x_22, 6); +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_22); +x_39 = l_Lean_Kernel_enableDiag(x_33, x_14); +x_40 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7; +x_41 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_34); +lean_ctor_set(x_41, 2, x_35); +lean_ctor_set(x_41, 3, x_36); +lean_ctor_set(x_41, 4, x_40); +lean_ctor_set(x_41, 5, x_37); +lean_ctor_set(x_41, 6, x_38); +x_42 = lean_st_ref_set(x_7, x_41, x_23); +x_43 = lean_ctor_get(x_42, 1); +lean_inc(x_43); +lean_dec(x_42); +x_44 = lean_box(0); +x_45 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1(x_12, x_14, x_1, x_2, x_3, x_4, x_5, x_44, x_6, x_7, x_43); +return x_45; +} +} +else +{ +lean_object* x_46; lean_object* x_47; +x_46 = lean_box(0); +x_47 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1(x_12, x_14, x_1, x_2, x_3, x_4, x_5, x_46, x_6, x_7, x_17); +return x_47; +} +} } } LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4) { @@ -30594,6 +30855,17 @@ x_6 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_ToDepElimPattern_main___spec__1 return x_6; } } +LEAN_EXPORT lean_object* l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_8); +return x_13; +} +} LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_ToDepElimPattern_main___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -43931,66 +44203,6 @@ goto _start; } } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_3; -x_3 = 1; -return x_3; -} -else -{ -uint8_t x_4; -x_4 = 0; -return x_4; -} -} -else -{ -if (lean_obj_tag(x_2) == 0) -{ -uint8_t x_5; -x_5 = 0; -return x_5; -} -else -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_unbox(x_6); -if (x_7 == 0) -{ -lean_object* x_8; uint8_t x_9; -x_8 = lean_ctor_get(x_2, 0); -x_9 = lean_unbox(x_8); -if (x_9 == 0) -{ -uint8_t x_10; -x_10 = 1; -return x_10; -} -else -{ -uint8_t x_11; -x_11 = 0; -return x_11; -} -} -else -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_unbox(x_12); -return x_13; -} -} -} -} -} LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___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, lean_object* x_12, lean_object* x_13) { _start: { @@ -45076,7 +45288,7 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); x_15 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_getMatchGeneralizing_x3f___lambda__2___closed__1; -x_16 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13(x_1, x_15); +x_16 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_PrettyPrinter_delabCore___spec__4(x_1, x_15); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; @@ -45269,17 +45481,6 @@ x_6 = l_Array_mapMUnsafe_map___at___private_Lean_Elab_Match_0__Lean_Elab_Term_el return x_6; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13___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_Elab_Match_0__Lean_Elab_Term_elabMatchAux___spec__13(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_Elab_Match_0__Lean_Elab_Term_elabMatchAux___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: { @@ -50896,10 +51097,22 @@ l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternIn lean_mark_persistent(l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___spec__4___rarg___closed__3); l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1___closed__1 = _init_l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_ToDepElimPattern_savePatternInfo_go___lambda__1___closed__1); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___lambda__1___closed__1); l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__1 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__1(); lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__1); l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2(); lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__2); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__3); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__4); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__5); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__6); +l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7 = _init_l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7(); +lean_mark_persistent(l_Lean_withInPattern___at_Lean_Elab_Term_ToDepElimPattern_main___spec__3___closed__7); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPattern_main___spec__5___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 61f70542077b..47ed8fe74053 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -66,6 +66,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_regis LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_checkKinds___lambda__1___closed__2; @@ -141,6 +142,7 @@ lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_elabCommand___spec_ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabMutualDef___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___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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; @@ -259,6 +261,7 @@ uint8_t l_List_any___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___lambda__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_MutualClosure_Replacement_apply(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_isMultiConstant_x3f___spec__2(lean_object*, size_t, size_t); @@ -295,6 +298,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualD lean_object* l_Lean_FVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__2___closed__3; static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6; +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); 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___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__4; @@ -306,6 +310,8 @@ lean_object* l_Lean_Elab_Term_expandLetEqnsDecl(lean_object*, uint8_t, lean_obje static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__6___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, size_t, lean_object*); 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*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__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_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; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__2; @@ -404,6 +410,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWher LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__7; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabMutualDef_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___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_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*); @@ -449,6 +456,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___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*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__17; lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -516,7 +524,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__1___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__2; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__2___closed__4; @@ -532,6 +539,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_checkForH LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerminationHint___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5(lean_object*); +lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__6; uint8_t l_Lean_Elab_DefView_isInstance(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -571,6 +579,7 @@ static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___close 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*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_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*); @@ -587,6 +596,7 @@ lean_object* l_Lean_Meta_Closure_mkForall(lean_object*, lean_object*); lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Elab_PreDefinition_WF_Fix_0__Lean_Elab_WF_processSumCasesOn___spec__1___rarg___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_Elab_instInhabitedDefView; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendindMVarErrorMessage___closed__3; +extern lean_object* l_Lean_diagnostics; uint8_t l_Lean_Syntax_isIdent(lean_object*); lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -620,7 +630,6 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getPendind LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_MutualClosure_getModifiersForLetRecs___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__5___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_getUsedFVarsMap___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__7___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkLetRecClosures___spec__8(lean_object*, lean_object*, size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Closure_mkLambda(lean_object*, lean_object*); size_t lean_usize_mod(size_t, size_t); @@ -681,8 +690,8 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at_Lean_Elab_Command_elabMutualDef___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___spec__5___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12; 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_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*); @@ -742,6 +751,7 @@ static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_MutualDe LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabMutualDef(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofLevel(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_withFunLocalDecls_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -759,6 +769,7 @@ static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_registerFa static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_FixPoint_updateUsedVarsOf___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFunName___closed__2; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkClosureForAux___closed__1; lean_object* l_Lean_indentExpr(lean_object*); @@ -768,11 +779,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_getFu LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___closed__5; static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__3; uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); @@ -812,6 +821,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Mutua LEAN_EXPORT lean_object* l_Lean_getDocStringText___at_Lean_Elab_Command_elabMutualDef___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__4; static lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashable; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at_Lean_Elab_Command_elabMutualDef___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_declValToTerm___closed__1; @@ -836,6 +846,7 @@ LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at_Lean_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_elabMutualDef_go___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkInitialUsedFVarsMap___spec__3___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_expandWhereStructInst___lambda__2___closed__3; +static lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabMutualDef___spec__6(lean_object*, size_t, size_t, 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*); @@ -888,6 +899,7 @@ size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_getMVars(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__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__5___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_MutualClosure_mkFreeVarMap___spec__6___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_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*); @@ -24006,28 +24018,7 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3) { -_start: -{ -lean_object* x_4; uint8_t x_5; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -x_5 = l_Lean_KVMap_contains(x_1, x_4); -lean_dec(x_4); -if (x_5 == 0) -{ -lean_object* x_6; -x_6 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_1, x_2, x_3); -return x_6; -} -else -{ -lean_dec(x_2); -return x_1; -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_Term_checkForHiddenUnivLevels_visitLevel___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; @@ -24065,31 +24056,132 @@ return x_18; } } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(lean_object* x_1, lean_object* x_2, uint8_t x_3) { _start: { -lean_object* x_1; -x_1 = l_Lean_pp_universes; +lean_object* x_4; uint8_t x_5; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +x_5 = l_Lean_KVMap_contains(x_1, x_4); +lean_dec(x_4); +if (x_5 == 0) +{ +lean_object* x_6; +x_6 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_1, x_2, x_3); +return x_6; +} +else +{ +lean_dec(x_2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2() { +} +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_pp_funBinderTypes; +x_1 = l_Lean_maxRecDepth; return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__3() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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) { _start: { -lean_object* x_1; -x_1 = l_Lean_pp_letVarTypes; -return x_1; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; +x_10 = lean_ctor_get(x_7, 0); +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 3); +x_13 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1; +x_14 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_13); +x_15 = lean_ctor_get(x_7, 5); +x_16 = lean_ctor_get(x_7, 6); +x_17 = lean_ctor_get(x_7, 7); +x_18 = lean_ctor_get(x_7, 8); +x_19 = lean_ctor_get(x_7, 9); +x_20 = lean_ctor_get(x_7, 10); +x_21 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +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_12); +lean_inc(x_11); +lean_inc(x_10); +x_22 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_22, 0, x_10); +lean_ctor_set(x_22, 1, x_11); +lean_ctor_set(x_22, 2, x_1); +lean_ctor_set(x_22, 3, x_12); +lean_ctor_set(x_22, 4, x_14); +lean_ctor_set(x_22, 5, x_15); +lean_ctor_set(x_22, 6, x_16); +lean_ctor_set(x_22, 7, x_17); +lean_ctor_set(x_22, 8, x_18); +lean_ctor_set(x_22, 9, x_19); +lean_ctor_set(x_22, 10, x_20); +lean_ctor_set_uint8(x_22, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_22, sizeof(void*)*11 + 1, x_2); +x_23 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_3, x_4, x_5, x_22, x_8, x_9); +lean_dec(x_22); +return x_23; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__4() { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_10 = lean_ctor_get(x_7, 0); +x_11 = lean_ctor_get(x_7, 1); +x_12 = lean_ctor_get(x_7, 3); +x_13 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1; +x_14 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_13); +x_15 = lean_ctor_get(x_7, 5); +x_16 = lean_ctor_get(x_7, 6); +x_17 = lean_ctor_get(x_7, 7); +x_18 = lean_ctor_get(x_7, 8); +x_19 = lean_ctor_get(x_7, 9); +x_20 = lean_ctor_get(x_7, 10); +x_21 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +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_12); +lean_inc(x_11); +lean_inc(x_10); +x_22 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_22, 0, x_10); +lean_ctor_set(x_22, 1, x_11); +lean_ctor_set(x_22, 2, x_1); +lean_ctor_set(x_22, 3, x_12); +lean_ctor_set(x_22, 4, x_14); +lean_ctor_set(x_22, 5, x_15); +lean_ctor_set(x_22, 6, x_16); +lean_ctor_set(x_22, 7, x_17); +lean_ctor_set(x_22, 8, x_18); +lean_ctor_set(x_22, 9, x_19); +lean_ctor_set(x_22, 10, x_20); +lean_ctor_set_uint8(x_22, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_22, sizeof(void*)*11 + 1, x_2); +x_23 = l_Lean_indentExpr(x_3); +x_24 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___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 = 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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_26, x_4, x_5, x_22, x_8, x_9); +lean_dec(x_22); +return x_27; +} +} +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1() { _start: { lean_object* x_1; @@ -24097,16 +24189,16 @@ x_1 = lean_mk_string_from_bytes("invalid occurrence of universe level '", 38); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__5() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__4; +x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__3() { _start: { lean_object* x_1; @@ -24114,16 +24206,16 @@ x_1 = lean_mk_string_from_bytes("' at '", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__7() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6; +x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__8() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__5() { _start: { lean_object* x_1; @@ -24131,16 +24223,16 @@ x_1 = lean_mk_string_from_bytes("', it does not occur at the declaration type, n return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__8; +x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__7() { _start: { lean_object* x_1; @@ -24148,15 +24240,47 @@ x_1 = lean_mk_string_from_bytes("\nat declaration body", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11() { +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10; +x_1 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__7; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_funBinderTypes; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_letVarTypes; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_universes; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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: { @@ -24306,153 +24430,297 @@ lean_dec(x_1); lean_dec(x_33); if (x_36 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_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; uint8_t x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_37 = lean_ctor_get(x_10, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_10, 1); -lean_inc(x_38); -x_39 = lean_ctor_get(x_10, 2); -lean_inc(x_39); -x_40 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1; -x_41 = 1; +lean_object* x_37; lean_object* x_38; lean_object* x_103; lean_object* x_104; uint8_t 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; uint8_t x_113; uint8_t x_114; +x_103 = lean_ctor_get(x_10, 2); +lean_inc(x_103); +x_104 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12; +x_105 = 1; +x_106 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_103, x_104, x_105); +x_107 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11; +x_108 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_106, x_107); +x_109 = lean_st_ref_get(x_11, x_12); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = lean_ctor_get(x_110, 0); +lean_inc(x_112); +lean_dec(x_110); +x_113 = l_Lean_Kernel_isDiagnosticsEnabled(x_112); +lean_dec(x_112); +if (x_113 == 0) +{ +if (x_108 == 0) +{ +x_114 = x_105; +goto block_137; +} +else +{ +uint8_t x_138; +x_138 = 0; +x_114 = x_138; +goto block_137; +} +} +else +{ +if (x_108 == 0) +{ +uint8_t x_139; +x_139 = 0; +x_114 = x_139; +goto block_137; +} +else +{ +x_114 = x_105; +goto block_137; +} +} +block_102: +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; uint8_t x_76; +x_39 = lean_ctor_get(x_3, 5); lean_inc(x_39); -x_42 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_39, x_40, x_41); -x_43 = lean_ctor_get(x_10, 3); -lean_inc(x_43); -x_44 = lean_ctor_get(x_10, 4); -lean_inc(x_44); -x_45 = lean_ctor_get(x_10, 5); -lean_inc(x_45); -x_46 = lean_ctor_get(x_10, 6); -lean_inc(x_46); -x_47 = lean_ctor_get(x_10, 7); -lean_inc(x_47); -x_48 = lean_ctor_get(x_10, 8); -lean_inc(x_48); -x_49 = lean_ctor_get(x_10, 9); -lean_inc(x_49); -x_50 = lean_ctor_get(x_10, 10); -lean_inc(x_50); -x_51 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -x_52 = lean_ctor_get_uint8(x_10, sizeof(void*)*11 + 1); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_inc(x_45); -lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_38); -lean_inc(x_37); -x_53 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_53, 0, x_37); -lean_ctor_set(x_53, 1, x_38); -lean_ctor_set(x_53, 2, x_42); -lean_ctor_set(x_53, 3, x_43); -lean_ctor_set(x_53, 4, x_44); -lean_ctor_set(x_53, 5, x_45); -lean_ctor_set(x_53, 6, x_46); -lean_ctor_set(x_53, 7, x_47); -lean_ctor_set(x_53, 8, x_48); -lean_ctor_set(x_53, 9, x_49); -lean_ctor_set(x_53, 10, x_50); -lean_ctor_set_uint8(x_53, sizeof(void*)*11, x_51); -lean_ctor_set_uint8(x_53, sizeof(void*)*11 + 1, x_52); -lean_inc(x_5); -x_54 = l_Lean_indentExpr(x_5); -x_55 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__4; -x_56 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -x_57 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_55); -x_58 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_57, x_8, x_9, x_53, x_11, x_12); -lean_dec(x_53); -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 = lean_ctor_get(x_3, 5); -lean_inc(x_61); -x_62 = l_Lean_indentExpr(x_61); -x_63 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_63, 0, x_55); -lean_ctor_set(x_63, 1, x_62); -x_64 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_55); -x_65 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2; -x_66 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(x_39, x_65, x_41); -x_67 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__3; -x_68 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(x_66, x_67, x_41); -x_69 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_69, 0, x_37); -lean_ctor_set(x_69, 1, x_38); -lean_ctor_set(x_69, 2, x_68); -lean_ctor_set(x_69, 3, x_43); -lean_ctor_set(x_69, 4, x_44); -lean_ctor_set(x_69, 5, x_45); -lean_ctor_set(x_69, 6, x_46); -lean_ctor_set(x_69, 7, x_47); -lean_ctor_set(x_69, 8, x_48); -lean_ctor_set(x_69, 9, x_49); -lean_ctor_set(x_69, 10, x_50); -lean_ctor_set_uint8(x_69, sizeof(void*)*11, x_51); -lean_ctor_set_uint8(x_69, sizeof(void*)*11 + 1, x_52); -x_70 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_64, x_8, x_9, x_69, x_11, x_60); -lean_dec(x_69); -x_71 = lean_ctor_get(x_70, 0); -lean_inc(x_71); -x_72 = lean_ctor_get(x_70, 1); +x_40 = l_Lean_indentExpr(x_39); +x_41 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_check___lambda__2___closed__4; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_63 = lean_ctor_get(x_10, 2); +lean_inc(x_63); +x_64 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9; +x_65 = 1; +x_66 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(x_63, x_64, x_65); +x_67 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10; +x_68 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(x_66, x_67, x_65); +x_69 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11; +x_70 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_68, x_69); +x_71 = lean_st_ref_get(x_11, x_38); +x_72 = lean_ctor_get(x_71, 0); lean_inc(x_72); -lean_dec(x_70); -x_73 = l_Lean_MessageData_ofLevel(x_4); -x_74 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__5; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -x_76 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__7; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_ctor_get(x_3, 3); -lean_inc(x_78); +x_73 = lean_ctor_get(x_71, 1); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_ctor_get(x_72, 0); +lean_inc(x_74); +lean_dec(x_72); +x_75 = l_Lean_Kernel_isDiagnosticsEnabled(x_74); +lean_dec(x_74); +if (x_75 == 0) +{ +if (x_70 == 0) +{ +x_76 = x_65; +goto block_99; +} +else +{ +uint8_t x_100; +x_100 = 0; +x_76 = x_100; +goto block_99; +} +} +else +{ +if (x_70 == 0) +{ +uint8_t x_101; +x_101 = 0; +x_76 = x_101; +goto block_99; +} +else +{ +x_76 = x_65; +goto block_99; +} +} +block_62: +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_46 = l_Lean_MessageData_ofLevel(x_4); +x_47 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2; +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_Term_checkForHiddenUnivLevels_visitLevel___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 = lean_ctor_get(x_3, 3); +lean_inc(x_51); lean_dec(x_3); -x_79 = l_Lean_MessageData_ofName(x_78); -x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__9; -x_82 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_82, 0, x_80); -lean_ctor_set(x_82, 1, x_81); -x_83 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_59); -x_84 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11; -x_85 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_85, 0, x_83); -lean_ctor_set(x_85, 1, x_84); -x_86 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_71); -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_55); -x_88 = l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_72); +x_52 = l_Lean_MessageData_ofName(x_51); +x_53 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__6; +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_37); +x_57 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__8; +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 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_44); +x_60 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_41); +x_61 = l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(x_60, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_45); lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_88; +return x_61; +} +block_99: +{ +if (x_76 == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_77 = lean_st_ref_take(x_11, x_73); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +lean_dec(x_77); +x_80 = lean_ctor_get(x_78, 0); +lean_inc(x_80); +x_81 = l_Lean_Kernel_enableDiag(x_80, x_70); +x_82 = lean_ctor_get(x_78, 1); +lean_inc(x_82); +x_83 = lean_ctor_get(x_78, 2); +lean_inc(x_83); +x_84 = lean_ctor_get(x_78, 3); +lean_inc(x_84); +x_85 = lean_ctor_get(x_78, 5); +lean_inc(x_85); +x_86 = lean_ctor_get(x_78, 6); +lean_inc(x_86); +lean_dec(x_78); +x_87 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__5; +x_88 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_88, 0, x_81); +lean_ctor_set(x_88, 1, x_82); +lean_ctor_set(x_88, 2, x_83); +lean_ctor_set(x_88, 3, x_84); +lean_ctor_set(x_88, 4, x_87); +lean_ctor_set(x_88, 5, x_85); +lean_ctor_set(x_88, 6, x_86); +x_89 = lean_st_ref_set(x_11, x_88, x_79); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); +x_91 = lean_box(0); +x_92 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1(x_68, x_70, x_43, x_8, x_9, x_91, x_10, x_11, x_90); +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_44 = x_93; +x_45 = x_94; +goto block_62; } else { -lean_object* x_89; lean_object* x_90; +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_95 = lean_box(0); +x_96 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1(x_68, x_70, x_43, x_8, x_9, x_95, x_10, x_11, x_73); +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_44 = x_97; +x_45 = x_98; +goto block_62; +} +} +} +block_137: +{ +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; +x_115 = lean_st_ref_take(x_11, x_111); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_ctor_get(x_116, 0); +lean_inc(x_118); +x_119 = l_Lean_Kernel_enableDiag(x_118, x_108); +x_120 = lean_ctor_get(x_116, 1); +lean_inc(x_120); +x_121 = lean_ctor_get(x_116, 2); +lean_inc(x_121); +x_122 = lean_ctor_get(x_116, 3); +lean_inc(x_122); +x_123 = lean_ctor_get(x_116, 5); +lean_inc(x_123); +x_124 = lean_ctor_get(x_116, 6); +lean_inc(x_124); +lean_dec(x_116); +x_125 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__5___closed__5; +x_126 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_126, 0, x_119); +lean_ctor_set(x_126, 1, x_120); +lean_ctor_set(x_126, 2, x_121); +lean_ctor_set(x_126, 3, x_122); +lean_ctor_set(x_126, 4, x_125); +lean_ctor_set(x_126, 5, x_123); +lean_ctor_set(x_126, 6, x_124); +x_127 = lean_st_ref_set(x_11, x_126, x_117); +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +lean_dec(x_127); +x_129 = lean_box(0); +lean_inc(x_5); +x_130 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__2(x_106, x_108, x_5, x_8, x_9, x_129, x_10, x_11, x_128); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +lean_dec(x_130); +x_37 = x_131; +x_38 = x_132; +goto block_102; +} +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_133 = lean_box(0); +lean_inc(x_5); +x_134 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__2(x_106, x_108, x_5, x_8, x_9, x_133, x_10, x_11, x_111); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_37 = x_135; +x_38 = x_136; +goto block_102; +} +} +} +else +{ +lean_object* x_140; lean_object* x_141; lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); @@ -24460,16 +24728,16 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_89 = lean_box(0); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_12); -return x_90; +x_140 = lean_box(0); +x_141 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_141, 0, x_140); +lean_ctor_set(x_141, 1, x_12); +return x_141; } } else { -lean_object* x_91; lean_object* x_92; +lean_object* x_142; lean_object* x_143; lean_dec(x_33); lean_dec(x_10); lean_dec(x_9); @@ -24479,16 +24747,16 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_91 = lean_box(0); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_12); -return x_92; +x_142 = lean_box(0); +x_143 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_143, 0, x_142); +lean_ctor_set(x_143, 1, x_12); +return x_143; } } default: { -lean_object* x_93; lean_object* x_94; +lean_object* x_144; lean_object* x_145; lean_dec(x_10); lean_dec(x_9); lean_dec(x_7); @@ -24498,38 +24766,68 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_93 = lean_box(0); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_12); -return x_94; +x_144 = lean_box(0); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_144); +lean_ctor_set(x_145, 1, x_12); +return x_145; } } } } -LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { uint8_t x_4; lean_object* x_5; x_4 = lean_unbox(x_3); lean_dec(x_3); -x_5 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__1(x_1, x_2, x_4); +x_5 = l_Lean_Option_setIfNotSet___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(x_1, x_2, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_Elab_Term_checkForHiddenUnivLevels_visitLevel___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_throwError___at_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); lean_dec(x_2); -return x_10; +x_11 = l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__2(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_11; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { @@ -34571,6 +34869,8 @@ l_Nat_foldM_loop___at_Lean_Elab_Term_MutualClosure_pushMain___spec__1___closed__ 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___lambda__1___closed__1 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___lambda__1___closed__1); l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__1); l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__2(); @@ -34593,6 +34893,8 @@ l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10 = _init_l_Lean lean_mark_persistent(l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__10); l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11(); lean_mark_persistent(l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__11); +l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12 = _init_l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_checkForHiddenUnivLevels_visitLevel___closed__12); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_checkForHiddenUnivLevels___spec__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c index f6a084753a7b..8057e2821546 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Eqns.c @@ -43,31 +43,34 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreD LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__13___boxed(lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfI(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_simpEqnType_collect___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10; static lean_object* l_Lean_Elab_Eqns_instInhabitedEqnInfoCore___closed__2; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__3___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_tryURefl___closed__3; +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13; +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__8; lean_object* l_Lean_Meta_Match_isNamedPattern___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13; +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__2___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__3___closed__1; static lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1___closed__4; static lean_object* l_Lean_Elab_Eqns_funext_x3f___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkEqnTypes_go___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_EXPORT lean_object* l_Lean_Elab_Eqns_expand___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -78,6 +81,7 @@ lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpMatch_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__4; lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT uint8_t l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__3___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__11___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -86,7 +90,6 @@ LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Pr LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_getType_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___at_Lean_Elab_Eqns_simpEqnType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); @@ -95,12 +98,13 @@ extern lean_object* l_Lean_ForEachExprWhere_initCache; lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__32___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__33(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ForEachExpr_visit___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___spec__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_isTracingEnabledFor___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_deltaRHS_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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*); -static lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__18(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -118,40 +122,47 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreD uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__3___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__13(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_instMonad(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15; +extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__17___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_findCore___at_Lean_Meta_removeUnused___spec__1(lean_object*, lean_object*); -lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___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_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__23(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_tryURefl___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_Eqns_tryURefl___closed__4; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__2(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3(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_Eqns_simpEqnType___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_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_tryURefl___closed__5; LEAN_EXPORT lean_object* l_Lean_commitWhenSome_x3f___at_Lean_Elab_Eqns_funext_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__15(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__1___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___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_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__17___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__28___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldrM___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__7; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_smartUnfolding; extern lean_object* l_instInhabitedPUnit; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); lean_object* l_Lean_Expr_replaceFVarId(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__30___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); -lean_object* l_Lean_HashMap_insert___at_Lean_ForEachExpr_visit___spec__3(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -168,6 +179,7 @@ size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__14___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__7(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_runST___rarg(lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__23(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___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_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___lambda__2___closed__2; @@ -189,6 +201,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__1___boxed(lean_o LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_isIrrelevant___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__6; LEAN_EXPORT uint8_t l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__1(lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__29(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*); @@ -203,6 +216,7 @@ LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at___private_Lean_Elab_PreD LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_quickCmp___boxed(lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__15___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__15___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__1___closed__2; @@ -213,7 +227,6 @@ lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_obj static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__5; static lean_object* l_Lean_Elab_Eqns_expandRHS_x3f___closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17; LEAN_EXPORT lean_object* l_Lean_LocalContext_foldrM___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__2___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__26(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__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_forInAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__28(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -231,11 +244,13 @@ lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lea lean_object* l_Lean_Meta_delta_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__11___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__12(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___boxed(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_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1; +lean_object* l_Lean_HashMap_insert___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___closed__3; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__22(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_deltaRHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -260,44 +275,51 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefi static double l_Lean_addTrace___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__28___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6; static lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__29___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__2; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__7(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_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_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(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_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_lhsDependsOn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof___closed__1; +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkUnfoldProof___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___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_Eqns_tryURefl___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___lambda__1___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_expand(uint8_t, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Split_splitMatch(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 uint8_t l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__7; static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__9; +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___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_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_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Eqns_mkEqnTypes_go___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___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_tryURefl___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__2(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11; lean_object* l_Lean_Meta_simpIfTarget(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_expandRHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Eqns_mkUnfoldProof___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7; static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__4; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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_Lean_Elab_Eqns_simpEqnType___spec__1___closed__4; @@ -305,21 +327,20 @@ static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5; lean_object* l_Lean_Meta_isMatcherAppCore_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__34___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_Eqns_removeUnusedEqnHypotheses_go___spec__11(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14; lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_diagnostics; lean_object* l_Lean_LocalDecl_fvarId(lean_object*); lean_object* l_Lean_Expr_FindExtImpl_findUnsafe_x3f(lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpIf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8; LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10; static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__30(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Eqns_deltaRHS_x3f___lambda__1(lean_object*, lean_object*); lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__21___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_Eqns_simpEqnType___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__1___boxed(lean_object*); @@ -336,13 +357,15 @@ static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Eqns_mkUnfoldProof___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___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_Elab_Eqns_splitMatch_x3f_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); extern lean_object* l_Lean_inheritedTraceOptions; static lean_object* l_Lean_Elab_Eqns_funext_x3f___lambda__1___closed__1; +lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__7(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Match_isNamedPattern_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Split_simpMatchTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_lhsDependsOn___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -350,21 +373,23 @@ double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__21___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__22(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_instInhabitedEqnInfoCore; static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___closed__1; -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___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_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__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_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Elab_Eqns_simpEqnType_collect___spec__4___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__19___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___closed__1; +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__9___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__10(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_simpEqnType_collect___spec__5(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isConstructorApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16; -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_isIrrelevant___lambda__1___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15; LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visited___at_Lean_Elab_Eqns_simpEqnType_collect___spec__4(lean_object*); static lean_object* l_Lean_Elab_Eqns_tryContradiction___closed__1; +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12; static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__23___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_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__14(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -372,12 +397,11 @@ lean_object* l_Lean_Meta_Match_unfoldNamedPattern___lambda__2___boxed(lean_objec LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4; static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__25(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__3___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkEqnTypes_go___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_brecOnSuffix; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__9___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__11; @@ -386,30 +410,29 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_sim lean_object* l_Lean_MVarId_isAssigned___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487_(lean_object*); lean_object* l_Lean_Meta_mkEqRefl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___closed__2; static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof___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*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__3___closed__1; static lean_object* l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__1; -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18; static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__3; static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_pushDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_funext_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2; lean_object* l_Lean_LocalDecl_type(lean_object*); lean_object* l_Lean_exprDependsOn___at___private_Lean_Meta_CongrTheorems_0__Lean_Meta_mkCast_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -422,8 +445,6 @@ static lean_object* l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__4; lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___closed__2; -static lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6; -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__24(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__23___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__24___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__9(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -438,11 +459,9 @@ static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__21(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Eqns_deltaLHS___lambda__1(lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch___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_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19; size_t lean_usize_sub(size_t, size_t); static lean_object* l_Lean_Elab_Eqns_mkUnfoldProof___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_whnfReducibleLHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -458,6 +477,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_E LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__9(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); +static lean_object* l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_splitMatch_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__3(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_Eqns_removeUnusedEqnHypotheses(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -481,8 +501,6 @@ lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___s lean_object* l_Lean_Meta_mkFreshExprSyntheticOpaqueMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7; -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___closed__1; static lean_object* l_Lean_Elab_Eqns_expandRHS_x3f___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_findMatchToSplit_x3f___lambda__3___closed__2; @@ -490,7 +508,6 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreD static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__2; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldProof_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_HashSetImp_contains___at_Lean_ForEachExprWhere_checked___spec__1(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_Eqns_funext_x3f___lambda__2___closed__1; @@ -503,6 +520,7 @@ lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint static lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___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_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1; static lean_object* l_Lean_Elab_Eqns_simpEqnType___lambda__1___closed__1; uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_isIrrelevant___lambda__1(lean_object*, lean_object*); @@ -510,21 +528,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_deltaLHS(lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_whnfAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFVar(lean_object*); +uint8_t l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2; lean_object* l_Lean_Meta_mkConstWithFreshMVarLevels(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_Elab_Eqns_tryURefl___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_consumeMData(lean_object*); static lean_object* l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__2; -static lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__16(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__17(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Eqns_tryURefl___closed__2; static lean_object* l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__1___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__19___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_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__3___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -533,7 +554,7 @@ LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_PreD static lean_object* l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldrMAux___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn___spec__15(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_Eqns_removeUnusedEqnHypotheses_go___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_Elab_Eqns_mkUnfoldEq___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2(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_mkFreshId___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__14___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__7___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Elab_Eqns_instInhabitedEqnInfoCore___closed__1() { @@ -2251,7 +2272,7 @@ if (x_5 == 0) { uint8_t x_6; lean_inc(x_4); -x_6 = l_Lean_HashSetImp_contains___at_Lean_ForEachExprWhere_checked___spec__1(x_1, x_4); +x_6 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_1, x_4); if (x_6 == 0) { lean_object* x_7; @@ -2686,7 +2707,7 @@ if (x_43 == 0) lean_object* x_44; lean_free_object(x_34); lean_dec(x_41); -x_44 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_33); +x_44 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_33); x_4 = x_44; x_9 = x_42; goto _start; @@ -2713,7 +2734,7 @@ else lean_object* x_47; lean_free_object(x_34); lean_dec(x_41); -x_47 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_33); +x_47 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_33); x_4 = x_47; x_9 = x_42; goto _start; @@ -2733,7 +2754,7 @@ if (x_51 == 0) { lean_object* x_52; lean_dec(x_49); -x_52 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_33); +x_52 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_33); x_4 = x_52; x_9 = x_50; goto _start; @@ -2763,7 +2784,7 @@ else { lean_object* x_56; lean_dec(x_49); -x_56 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_33); +x_56 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_33); x_4 = x_56; x_9 = x_50; goto _start; @@ -2841,7 +2862,7 @@ if (x_68 == 0) lean_object* x_69; lean_dec(x_67); lean_dec(x_65); -x_69 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_58); +x_69 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_58); x_4 = x_69; x_9 = x_66; goto _start; @@ -2876,7 +2897,7 @@ else lean_object* x_73; lean_dec(x_67); lean_dec(x_65); -x_73 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_4, x_58); +x_73 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_4, x_58); x_4 = x_73; x_9 = x_66; goto _start; @@ -3073,143 +3094,138 @@ lean_dec(x_1); return x_9; } } -LEAN_EXPORT lean_object* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3) { -_start: -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); -x_5 = lean_alloc_ctor(1, 0, 1); -lean_ctor_set_uint8(x_5, 0, x_3); -x_6 = l_Lean_KVMap_insertCore(x_1, x_4, x_5); -return x_6; -} -} -static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___closed__1() { +static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_smartUnfolding; +x_1 = l_Lean_maxRecDepth; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_tryURefl(lean_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_Eqns_tryURefl___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) { _start: { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_8 = lean_ctor_get(x_4, 2); -x_9 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); -x_10 = l_Lean_Elab_Eqns_tryURefl___closed__1; -x_11 = 0; -x_12 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_8, x_10, x_11); -lean_ctor_set(x_4, 2, x_12); -x_13 = l_Lean_MVarId_refl(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_13) == 0) +uint8_t 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_uint8(x_7, sizeof(void*)*11); +x_12 = lean_ctor_get(x_7, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_7, 2); +lean_dec(x_13); +x_14 = l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_ctor_set(x_7, 4, x_15); +lean_ctor_set(x_7, 2, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*11 + 1, x_2); +x_16 = l_Lean_MVarId_refl(x_3, x_4, x_5, x_7, x_8, x_9); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_15; uint8_t x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -x_16 = 1; -x_17 = lean_box(x_16); -lean_ctor_set(x_13, 0, x_17); -return x_13; +lean_object* x_18; uint8_t x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +x_19 = 1; +x_20 = lean_box(x_19); +lean_ctor_set(x_16, 0, x_20); +return x_16; } else { -lean_object* x_18; uint8_t x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = 1; -x_20 = lean_box(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_18); -return x_21; +lean_object* x_21; uint8_t x_22; lean_object* x_23; lean_object* x_24; +x_21 = lean_ctor_get(x_16, 1); +lean_inc(x_21); +lean_dec(x_16); +x_22 = 1; +x_23 = lean_box(x_22); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_21); +return x_24; } } else { -uint8_t x_22; -x_22 = !lean_is_exclusive(x_13); -if (x_22 == 0) +uint8_t x_25; +x_25 = !lean_is_exclusive(x_16); +if (x_25 == 0) { -lean_object* x_23; uint8_t x_24; -x_23 = lean_ctor_get(x_13, 0); -x_24 = l_Lean_Exception_isRuntime(x_23); -if (x_24 == 0) +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_16, 0); +x_27 = l_Lean_Exception_isRuntime(x_26); +if (x_27 == 0) { -lean_object* x_25; -lean_dec(x_23); -x_25 = lean_box(x_11); -lean_ctor_set_tag(x_13, 0); -lean_ctor_set(x_13, 0, x_25); -return x_13; +uint8_t x_28; lean_object* x_29; +lean_dec(x_26); +x_28 = 0; +x_29 = lean_box(x_28); +lean_ctor_set_tag(x_16, 0); +lean_ctor_set(x_16, 0, x_29); +return x_16; } else { -if (x_9 == 0) +if (x_11 == 0) { -return x_13; +return x_16; } else { -lean_object* x_26; -lean_dec(x_23); -x_26 = lean_box(x_11); -lean_ctor_set_tag(x_13, 0); -lean_ctor_set(x_13, 0, x_26); -return x_13; +uint8_t x_30; lean_object* x_31; +lean_dec(x_26); +x_30 = 0; +x_31 = lean_box(x_30); +lean_ctor_set_tag(x_16, 0); +lean_ctor_set(x_16, 0, x_31); +return x_16; } } } else { -lean_object* x_27; lean_object* x_28; uint8_t 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_dec(x_13); -x_29 = l_Lean_Exception_isRuntime(x_27); -if (x_29 == 0) +lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_32 = lean_ctor_get(x_16, 0); +x_33 = lean_ctor_get(x_16, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_16); +x_34 = l_Lean_Exception_isRuntime(x_32); +if (x_34 == 0) { -lean_object* x_30; lean_object* x_31; -lean_dec(x_27); -x_30 = lean_box(x_11); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -return x_31; +uint8_t x_35; lean_object* x_36; lean_object* x_37; +lean_dec(x_32); +x_35 = 0; +x_36 = lean_box(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_33); +return x_37; } else { -if (x_9 == 0) +if (x_11 == 0) { -lean_object* x_32; -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_27); -lean_ctor_set(x_32, 1, x_28); -return x_32; +lean_object* x_38; +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_32); +lean_ctor_set(x_38, 1, x_33); +return x_38; } else { -lean_object* x_33; lean_object* x_34; -lean_dec(x_27); -x_33 = lean_box(x_11); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_28); -return x_34; +uint8_t x_39; lean_object* x_40; lean_object* x_41; +lean_dec(x_32); +x_39 = 0; +x_40 = lean_box(x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_33); +return x_41; } } } @@ -3217,147 +3233,328 @@ return x_34; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; uint8_t x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_35 = lean_ctor_get(x_4, 0); -x_36 = lean_ctor_get(x_4, 1); -x_37 = lean_ctor_get(x_4, 2); -x_38 = lean_ctor_get(x_4, 3); -x_39 = lean_ctor_get(x_4, 4); -x_40 = lean_ctor_get(x_4, 5); -x_41 = lean_ctor_get(x_4, 6); -x_42 = lean_ctor_get(x_4, 7); -x_43 = lean_ctor_get(x_4, 8); -x_44 = lean_ctor_get(x_4, 9); -x_45 = lean_ctor_get(x_4, 10); -x_46 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); -x_47 = lean_ctor_get_uint8(x_4, sizeof(void*)*11 + 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_42 = lean_ctor_get(x_7, 0); +x_43 = lean_ctor_get(x_7, 1); +x_44 = lean_ctor_get(x_7, 3); +x_45 = lean_ctor_get(x_7, 5); +x_46 = lean_ctor_get(x_7, 6); +x_47 = lean_ctor_get(x_7, 7); +x_48 = lean_ctor_get(x_7, 8); +x_49 = lean_ctor_get(x_7, 9); +x_50 = lean_ctor_get(x_7, 10); +x_51 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); lean_inc(x_45); lean_inc(x_44); lean_inc(x_43); lean_inc(x_42); -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_4); -x_48 = l_Lean_Elab_Eqns_tryURefl___closed__1; -x_49 = 0; -x_50 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_37, x_48, x_49); -x_51 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_51, 0, x_35); -lean_ctor_set(x_51, 1, x_36); -lean_ctor_set(x_51, 2, x_50); -lean_ctor_set(x_51, 3, x_38); -lean_ctor_set(x_51, 4, x_39); -lean_ctor_set(x_51, 5, x_40); -lean_ctor_set(x_51, 6, x_41); -lean_ctor_set(x_51, 7, x_42); -lean_ctor_set(x_51, 8, x_43); -lean_ctor_set(x_51, 9, x_44); -lean_ctor_set(x_51, 10, x_45); -lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_46); -lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_47); -x_52 = l_Lean_MVarId_refl(x_1, x_2, x_3, x_51, x_5, x_6); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; lean_object* x_57; -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; +lean_dec(x_7); +x_52 = l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1; +x_53 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_52); +x_54 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_54, 0, x_42); +lean_ctor_set(x_54, 1, x_43); +lean_ctor_set(x_54, 2, x_1); +lean_ctor_set(x_54, 3, x_44); +lean_ctor_set(x_54, 4, x_53); +lean_ctor_set(x_54, 5, x_45); +lean_ctor_set(x_54, 6, x_46); +lean_ctor_set(x_54, 7, x_47); +lean_ctor_set(x_54, 8, x_48); +lean_ctor_set(x_54, 9, x_49); +lean_ctor_set(x_54, 10, x_50); +lean_ctor_set_uint8(x_54, sizeof(void*)*11, x_51); +lean_ctor_set_uint8(x_54, sizeof(void*)*11 + 1, x_2); +x_55 = l_Lean_MVarId_refl(x_3, x_4, x_5, x_54, x_8, x_9); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; +x_56 = lean_ctor_get(x_55, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_57 = x_55; } else { - lean_dec_ref(x_52); - x_54 = lean_box(0); + lean_dec_ref(x_55); + x_57 = lean_box(0); } -x_55 = 1; -x_56 = lean_box(x_55); -if (lean_is_scalar(x_54)) { - x_57 = lean_alloc_ctor(0, 2, 0); +x_58 = 1; +x_59 = lean_box(x_58); +if (lean_is_scalar(x_57)) { + x_60 = lean_alloc_ctor(0, 2, 0); } else { - x_57 = x_54; + x_60 = x_57; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_53); -return x_57; +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_56); +return x_60; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_58 = lean_ctor_get(x_52, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_52, 1); -lean_inc(x_59); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_60 = x_52; +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_61 = lean_ctor_get(x_55, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_55, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_63 = x_55; } else { - lean_dec_ref(x_52); - x_60 = lean_box(0); + lean_dec_ref(x_55); + x_63 = lean_box(0); } -x_61 = l_Lean_Exception_isRuntime(x_58); -if (x_61 == 0) +x_64 = l_Lean_Exception_isRuntime(x_61); +if (x_64 == 0) { -lean_object* x_62; lean_object* x_63; -lean_dec(x_58); -x_62 = lean_box(x_49); -if (lean_is_scalar(x_60)) { - x_63 = lean_alloc_ctor(0, 2, 0); +uint8_t x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_61); +x_65 = 0; +x_66 = lean_box(x_65); +if (lean_is_scalar(x_63)) { + x_67 = lean_alloc_ctor(0, 2, 0); } else { - x_63 = x_60; - lean_ctor_set_tag(x_63, 0); + x_67 = x_63; + lean_ctor_set_tag(x_67, 0); } -lean_ctor_set(x_63, 0, x_62); -lean_ctor_set(x_63, 1, x_59); -return x_63; +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_62); +return x_67; } else { -if (x_46 == 0) +if (x_51 == 0) { -lean_object* x_64; -if (lean_is_scalar(x_60)) { - x_64 = lean_alloc_ctor(1, 2, 0); +lean_object* x_68; +if (lean_is_scalar(x_63)) { + x_68 = lean_alloc_ctor(1, 2, 0); } else { - x_64 = x_60; + x_68 = x_63; } -lean_ctor_set(x_64, 0, x_58); -lean_ctor_set(x_64, 1, x_59); -return x_64; +lean_ctor_set(x_68, 0, x_61); +lean_ctor_set(x_68, 1, x_62); +return x_68; } else { -lean_object* x_65; lean_object* x_66; -lean_dec(x_58); -x_65 = lean_box(x_49); -if (lean_is_scalar(x_60)) { - x_66 = lean_alloc_ctor(0, 2, 0); +uint8_t x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_61); +x_69 = 0; +x_70 = lean_box(x_69); +if (lean_is_scalar(x_63)) { + x_71 = lean_alloc_ctor(0, 2, 0); } else { - x_66 = x_60; - lean_ctor_set_tag(x_66, 0); + x_71 = x_63; + lean_ctor_set_tag(x_71, 0); +} +lean_ctor_set(x_71, 0, x_70); +lean_ctor_set(x_71, 1, x_62); +return x_71; } -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_59); -return x_66; } } } } } +static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_smartUnfolding; +return x_1; +} } -LEAN_EXPORT lean_object* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___closed__2() { _start: { -uint8_t x_4; lean_object* x_5; -x_4 = lean_unbox(x_3); -lean_dec(x_3); -x_5 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_1, x_2, x_4); -return x_5; +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___closed__3() { +_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_Elab_Eqns_tryURefl___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Eqns_tryURefl___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_Elab_Eqns_tryURefl___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Eqns_tryURefl___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; +} +} +static lean_object* _init_l_Lean_Elab_Eqns_tryURefl___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Eqns_tryURefl___closed__5; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_tryURefl(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; lean_object* x_11; uint8_t 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; +x_7 = lean_ctor_get(x_4, 2); +lean_inc(x_7); +x_8 = l_Lean_Elab_Eqns_tryURefl___closed__1; +x_9 = 0; +x_10 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_7, x_8, x_9); +x_11 = l_Lean_Elab_Eqns_tryURefl___closed__2; +x_12 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_11); +x_13 = lean_st_ref_get(x_5, x_6); +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_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Kernel_isDiagnosticsEnabled(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +if (x_12 == 0) +{ +uint8_t x_47; +x_47 = 1; +x_18 = x_47; +goto block_46; +} +else +{ +x_18 = x_9; +goto block_46; +} +} +else +{ +if (x_12 == 0) +{ +x_18 = x_9; +goto block_46; +} +else +{ +uint8_t x_48; +x_48 = 1; +x_18 = x_48; +goto block_46; +} +} +block_46: +{ +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_st_ref_take(x_5, x_15); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = !lean_is_exclusive(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_20, 0); +x_24 = lean_ctor_get(x_20, 4); +lean_dec(x_24); +x_25 = l_Lean_Kernel_enableDiag(x_23, x_12); +x_26 = l_Lean_Elab_Eqns_tryURefl___closed__6; +lean_ctor_set(x_20, 4, x_26); +lean_ctor_set(x_20, 0, x_25); +x_27 = lean_st_ref_set(x_5, x_20, x_21); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Eqns_tryURefl___lambda__1(x_10, x_12, x_1, x_2, x_3, x_29, x_4, x_5, 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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +x_33 = lean_ctor_get(x_20, 2); +x_34 = lean_ctor_get(x_20, 3); +x_35 = lean_ctor_get(x_20, 5); +x_36 = lean_ctor_get(x_20, 6); +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_20); +x_37 = l_Lean_Kernel_enableDiag(x_31, x_12); +x_38 = l_Lean_Elab_Eqns_tryURefl___closed__6; +x_39 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_32); +lean_ctor_set(x_39, 2, x_33); +lean_ctor_set(x_39, 3, x_34); +lean_ctor_set(x_39, 4, x_38); +lean_ctor_set(x_39, 5, x_35); +lean_ctor_set(x_39, 6, x_36); +x_40 = lean_st_ref_set(x_5, x_39, x_21); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_box(0); +x_43 = l_Lean_Elab_Eqns_tryURefl___lambda__1(x_10, x_12, x_1, x_2, x_3, x_42, x_4, x_5, x_41); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_box(0); +x_45 = l_Lean_Elab_Eqns_tryURefl___lambda__1(x_10, x_12, x_1, x_2, x_3, x_44, x_4, x_5, x_15); +return x_45; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_tryURefl___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_2); +lean_dec(x_2); +x_11 = l_Lean_Elab_Eqns_tryURefl___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_11; } } static lean_object* _init_l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__1() { @@ -3781,7 +3978,7 @@ x_8 = lean_ctor_get(x_6, 1); lean_inc(x_8); lean_dec(x_6); lean_inc(x_1); -x_9 = l_Lean_HashSetImp_contains___at_Lean_ForEachExprWhere_checked___spec__1(x_8, x_1); +x_9 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_8, x_1); if (x_9 == 0) { lean_object* x_10; @@ -3800,7 +3997,7 @@ if (x_13 == 0) { lean_object* x_14; lean_object* x_15; lean_object* x_16; x_14 = lean_ctor_get(x_11, 1); -x_15 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_14, x_1); +x_15 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_14, x_1); lean_ctor_set(x_11, 1, x_15); x_16 = lean_st_ref_set(x_2, x_11, x_12); if (lean_obj_tag(x_16) == 0) @@ -3862,7 +4059,7 @@ x_30 = lean_ctor_get(x_11, 1); lean_inc(x_30); lean_inc(x_29); lean_dec(x_11); -x_31 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_30, x_1); +x_31 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_30, x_1); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_29); lean_ctor_set(x_32, 1, x_31); @@ -3963,7 +4160,7 @@ x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); lean_dec(x_49); lean_inc(x_1); -x_52 = l_Lean_HashSetImp_contains___at_Lean_ForEachExprWhere_checked___spec__1(x_51, x_1); +x_52 = l_Lean_HashSetImp_contains___at_Lean_CollectLevelParams_visitExpr___spec__1(x_51, x_1); if (x_52 == 0) { lean_object* x_53; @@ -3988,7 +4185,7 @@ if (lean_is_exclusive(x_54)) { lean_dec_ref(x_54); x_58 = lean_box(0); } -x_59 = l_Lean_HashSetImp_insert___at_Lean_ForEachExprWhere_checked___spec__3(x_57, x_1); +x_59 = l_Lean_HashSetImp_insert___at_Lean_CollectLevelParams_visitExpr___spec__3(x_57, x_1); if (lean_is_scalar(x_58)) { x_60 = lean_alloc_ctor(0, 2, 0); } else { @@ -4118,398 +4315,392 @@ x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_checked___at_Lean_Elab_ return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___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_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___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) { _start: { -lean_dec(x_4); +lean_dec(x_5); switch (lean_obj_tag(x_1)) { case 5: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); +x_9 = lean_ctor_get(x_1, 1); +lean_inc(x_9); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_9 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_7, x_5, x_6); -if (lean_obj_tag(x_9) == 0) +x_10 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_8, x_6, x_7); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_9, 1); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_8, x_5, x_10); -return x_11; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_9, x_6, x_11); +return x_12; } else { -uint8_t x_12; -lean_dec(x_8); -lean_dec(x_5); +uint8_t x_13; +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_12 = !lean_is_exclusive(x_9); -if (x_12 == 0) +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) { -return x_9; +return x_10; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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_inc(x_13); -lean_dec(x_9); -x_15 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_15, 0, x_13); -lean_ctor_set(x_15, 1, x_14); -return x_15; +lean_dec(x_10); +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; } } } case 6: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_1, 1); -lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 2); +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_1, 1); lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 2); +lean_inc(x_18); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_18 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_16, x_5, x_6); -if (lean_obj_tag(x_18) == 0) +x_19 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_17, x_6, x_7); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_17, x_5, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_18, x_6, x_20); +return x_21; } else { -uint8_t x_21; -lean_dec(x_17); -lean_dec(x_5); +uint8_t x_22; +lean_dec(x_18); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) { -return x_18; +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 0); -x_23 = lean_ctor_get(x_18, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_18); -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_19); +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; } } } case 7: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_1, 1); -lean_inc(x_25); -x_26 = lean_ctor_get(x_1, 2); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_1, 1); lean_inc(x_26); +x_27 = lean_ctor_get(x_1, 2); +lean_inc(x_27); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_27 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_25, x_5, x_6); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_26, x_6, x_7); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; -x_28 = lean_ctor_get(x_27, 1); -lean_inc(x_28); -lean_dec(x_27); -x_29 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_26, x_5, x_28); -return x_29; +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_27, x_6, x_29); +return x_30; } else { -uint8_t x_30; -lean_dec(x_26); -lean_dec(x_5); +uint8_t x_31; +lean_dec(x_27); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_27); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -return x_27; +return x_28; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_27, 0); -x_32 = lean_ctor_get(x_27, 1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +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_inc(x_31); -lean_dec(x_27); -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_dec(x_28); +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; } } } case 8: { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_1, 1); -lean_inc(x_34); -x_35 = lean_ctor_get(x_1, 2); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_1, 1); lean_inc(x_35); -x_36 = lean_ctor_get(x_1, 3); +x_36 = lean_ctor_get(x_1, 2); lean_inc(x_36); +x_37 = lean_ctor_get(x_1, 3); +lean_inc(x_37); lean_dec(x_1); -lean_inc(x_5); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_37 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_34, x_5, x_6); -if (lean_obj_tag(x_37) == 0) +x_38 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_35, x_6, x_7); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -lean_inc(x_5); +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_38, 1); +lean_inc(x_39); +lean_dec(x_38); +lean_inc(x_6); lean_inc(x_3); lean_inc(x_2); -x_39 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_35, x_5, x_38); -if (lean_obj_tag(x_39) == 0) +x_40 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_36, x_6, x_39); +if (lean_obj_tag(x_40) == 0) { -lean_object* x_40; lean_object* x_41; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_36, x_5, x_40); -return x_41; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_37, x_6, x_41); +return x_42; } else { -uint8_t x_42; -lean_dec(x_36); -lean_dec(x_5); +uint8_t x_43; +lean_dec(x_37); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_42 = !lean_is_exclusive(x_39); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_40); +if (x_43 == 0) { -return x_39; +return x_40; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_39, 0); -x_44 = lean_ctor_get(x_39, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_40, 0); +x_45 = lean_ctor_get(x_40, 1); +lean_inc(x_45); lean_inc(x_44); -lean_inc(x_43); -lean_dec(x_39); -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_dec(x_40); +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_46; +uint8_t x_47; +lean_dec(x_37); lean_dec(x_36); -lean_dec(x_35); -lean_dec(x_5); +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_37); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_38); +if (x_47 == 0) { -return x_37; +return x_38; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_37, 0); -x_48 = lean_ctor_get(x_37, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_38, 0); +x_49 = lean_ctor_get(x_38, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_37); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_38); +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; } } } case 10: { -lean_object* x_50; lean_object* x_51; -x_50 = lean_ctor_get(x_1, 1); -lean_inc(x_50); +lean_object* x_51; lean_object* x_52; +x_51 = lean_ctor_get(x_1, 1); +lean_inc(x_51); lean_dec(x_1); -x_51 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_50, x_5, x_6); -return x_51; +x_52 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_51, x_6, x_7); +return x_52; } case 11: { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_1, 2); -lean_inc(x_52); +lean_object* x_53; lean_object* x_54; +x_53 = lean_ctor_get(x_1, 2); +lean_inc(x_53); lean_dec(x_1); -x_53 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_52, x_5, x_6); -return x_53; +x_54 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_2, x_3, x_4, x_53, x_6, x_7); +return x_54; } default: { -lean_object* x_54; lean_object* x_55; -lean_dec(x_5); +lean_object* x_55; lean_object* x_56; +lean_dec(x_6); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -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_6); -return x_55; +x_55 = lean_box(0); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_7); +return x_56; } } } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(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_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___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) { _start: { -lean_object* x_6; -lean_inc(x_3); -x_6 = l_Lean_ForEachExprWhere_visited___at_Lean_Elab_Eqns_simpEqnType_collect___spec__4___rarg(x_3, x_4, x_5); -if (lean_obj_tag(x_6) == 0) +lean_object* x_7; +lean_inc(x_4); +x_7 = l_Lean_ForEachExprWhere_visited___at_Lean_Elab_Eqns_simpEqnType_collect___spec__4___rarg(x_4, x_5, x_6); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_7; uint8_t x_8; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_unbox(x_7); -lean_dec(x_7); -if (x_8 == 0) +lean_object* x_8; uint8_t x_9; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_6, 1); -lean_inc(x_9); -lean_dec(x_6); +lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); lean_inc(x_1); -lean_inc(x_3); -x_10 = lean_apply_1(x_1, x_3); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) +lean_inc(x_4); +x_11 = lean_apply_1(x_1, x_4); +x_12 = lean_unbox(x_11); +lean_dec(x_11); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_3, x_1, x_2, x_12, x_4, x_9); -return x_13; +lean_object* x_13; lean_object* x_14; +x_13 = lean_box(0); +x_14 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_4, x_1, x_2, x_3, x_13, x_5, x_10); +return x_14; } else { -lean_object* x_14; -lean_inc(x_3); -x_14 = l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_simpEqnType_collect___spec__5___rarg(x_3, x_4, x_9); -if (lean_obj_tag(x_14) == 0) +lean_object* x_15; +lean_inc(x_4); +x_15 = l_Lean_ForEachExprWhere_checked___at_Lean_Elab_Eqns_simpEqnType_collect___spec__5___rarg(x_4, x_5, x_10); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; uint8_t x_16; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_unbox(x_15); +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_object* x_19; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); lean_dec(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); lean_inc(x_2); -lean_inc(x_3); -x_18 = lean_apply_2(x_2, x_3, x_17); -if (lean_obj_tag(x_18) == 0) +lean_inc(x_4); +x_19 = lean_apply_2(x_2, x_4, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +if (x_3 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_19, 1); lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_3, x_1, x_2, x_19, x_4, x_20); -return x_21; +lean_dec(x_19); +x_21 = lean_box(0); +x_22 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_4, x_1, x_2, x_3, x_21, x_5, x_20); +return x_22; } else { -uint8_t x_22; +uint8_t x_23; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_22 = !lean_is_exclusive(x_18); -if (x_22 == 0) -{ -return x_18; -} -else +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) { -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; -} -} +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_19, 0); +lean_dec(x_24); +x_25 = lean_box(0); +lean_ctor_set(x_19, 0, x_25); +return x_19; } else { lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_14, 1); +x_26 = lean_ctor_get(x_19, 1); lean_inc(x_26); -lean_dec(x_14); +lean_dec(x_19); x_27 = lean_box(0); -x_28 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_3, x_1, x_2, x_27, x_4, 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_26); return x_28; } } +} else { uint8_t x_29; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_14); +x_29 = !lean_is_exclusive(x_19); if (x_29 == 0) { -return x_14; +return x_19; } else { lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_14, 0); -x_31 = lean_ctor_get(x_14, 1); +x_30 = lean_ctor_get(x_19, 0); +x_31 = lean_ctor_get(x_19, 1); lean_inc(x_31); lean_inc(x_30); -lean_dec(x_14); +lean_dec(x_19); x_32 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); @@ -4517,62 +4708,100 @@ return x_32; } } } +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_15, 1); +lean_inc(x_33); +lean_dec(x_15); +x_34 = lean_box(0); +x_35 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_4, x_1, x_2, x_3, x_34, x_5, x_33); +return x_35; +} } else { -uint8_t x_33; +uint8_t x_36; +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_6); -if (x_33 == 0) +x_36 = !lean_is_exclusive(x_15); +if (x_36 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_6, 0); -lean_dec(x_34); -x_35 = lean_box(0); -lean_ctor_set(x_6, 0, x_35); -return x_6; +return x_15; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_6, 1); -lean_inc(x_36); -lean_dec(x_6); -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; +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_15, 0); +x_38 = lean_ctor_get(x_15, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_15); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} } } } else { -uint8_t x_39; +uint8_t x_40; +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_6); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_7); +if (x_40 == 0) { -return x_6; +lean_object* x_41; lean_object* x_42; +x_41 = lean_ctor_get(x_7, 0); +lean_dec(x_41); +x_42 = lean_box(0); +lean_ctor_set(x_7, 0, x_42); +return x_7; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_6, 0); -x_41 = lean_ctor_get(x_6, 1); -lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_6); -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_7, 1); +lean_inc(x_43); +lean_dec(x_7); +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 +{ +uint8_t x_46; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_46 = !lean_is_exclusive(x_7); +if (x_46 == 0) +{ +return x_7; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_7, 0); +x_48 = lean_ctor_get(x_7, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_7); +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; } } } @@ -4581,131 +4810,131 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_si _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg), 5, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___boxed), 6, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_ForEachExprWhere_initCache; -x_6 = lean_st_mk_ref(x_5, x_4); -if (lean_obj_tag(x_6) == 0) +lean_object* x_6; lean_object* x_7; +x_6 = l_Lean_ForEachExprWhere_initCache; +x_7 = lean_st_mk_ref(x_6, x_5); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); -lean_dec(x_6); -lean_inc(x_7); -x_9 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_1, x_2, x_3, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_8); +x_10 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_1, x_2, x_4, x_3, x_8, x_9); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -lean_dec(x_9); -x_12 = lean_st_ref_get(x_7, x_11); -lean_dec(x_7); -if (lean_obj_tag(x_12) == 0) +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_st_ref_get(x_8, x_12); +lean_dec(x_8); +if (lean_obj_tag(x_13) == 0) { -uint8_t x_13; -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_14; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -lean_ctor_set(x_12, 0, x_10); -return x_12; +lean_object* x_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_11); +return x_13; } else { -lean_object* x_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_10); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_11); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -uint8_t x_17; -lean_dec(x_10); -x_17 = !lean_is_exclusive(x_12); -if (x_17 == 0) +uint8_t x_18; +lean_dec(x_11); +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) { -return x_12; +return x_13; } 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_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_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_dec(x_13); +x_21 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +return x_21; } } } else { -uint8_t x_21; -lean_dec(x_7); -x_21 = !lean_is_exclusive(x_9); -if (x_21 == 0) +uint8_t x_22; +lean_dec(x_8); +x_22 = !lean_is_exclusive(x_10); +if (x_22 == 0) { -return x_9; +return x_10; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_9, 0); -x_23 = lean_ctor_get(x_9, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_10, 0); +x_24 = lean_ctor_get(x_10, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_9); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_10); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -uint8_t x_25; +uint8_t x_26; lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_6); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_7); +if (x_26 == 0) { -return x_6; +return x_7; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_6, 0); -x_27 = lean_ctor_get(x_6, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_7, 0); +x_28 = lean_ctor_get(x_7, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_6); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_dec(x_7); +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; } } } @@ -4714,7 +4943,7 @@ LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpE _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg), 4, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg___boxed), 5, 0); return x_2; } } @@ -4854,7 +5083,7 @@ x_4 = lean_box(0); x_5 = lean_st_mk_ref(x_4, x_3); 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_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); x_7 = lean_ctor_get(x_5, 1); @@ -4864,62 +5093,63 @@ lean_inc(x_6); x_8 = lean_alloc_closure((void*)(l_Lean_Elab_Eqns_simpEqnType_collect___lambda__1), 3, 1); lean_closure_set(x_8, 0, x_6); x_9 = l_Lean_Elab_Eqns_simpEqnType_collect___lambda__2___closed__1; -x_10 = l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(x_9, x_8, x_1, x_7); -if (lean_obj_tag(x_10) == 0) +x_10 = 0; +x_11 = l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(x_9, x_8, x_1, x_10, x_7); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = lean_st_ref_get(x_6, x_11); +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = lean_st_ref_get(x_6, x_12); lean_dec(x_6); -return x_12; +return x_13; } else { -uint8_t x_13; +uint8_t x_14; lean_dec(x_6); -x_13 = !lean_is_exclusive(x_10); -if (x_13 == 0) +x_14 = !lean_is_exclusive(x_11); +if (x_14 == 0) { -return x_10; +return x_11; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_10, 0); -x_15 = lean_ctor_get(x_10, 1); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_11, 0); +x_16 = lean_ctor_get(x_11, 1); +lean_inc(x_16); lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_10); -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_11); +x_17 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } } else { -uint8_t x_17; +uint8_t x_18; lean_dec(x_1); -x_17 = !lean_is_exclusive(x_5); -if (x_17 == 0) +x_18 = !lean_is_exclusive(x_5); +if (x_18 == 0) { return x_5; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_5, 0); -x_19 = lean_ctor_get(x_5, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_5, 0); +x_20 = lean_ctor_get(x_5, 1); +lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); lean_dec(x_5); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; +x_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; } } } @@ -4952,6 +5182,36 @@ lean_dec(x_2); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___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) { +_start: +{ +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_4); +lean_dec(x_4); +x_9 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg___lambda__1(x_1, x_2, x_3, x_8, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___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) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_ForEachExprWhere_visit_go___at_Lean_Elab_Eqns_simpEqnType_collect___spec__3___rarg(x_1, x_2, x_7, x_4, x_5, x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg___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_6; lean_object* x_7; +x_6 = lean_unbox(x_4); +lean_dec(x_4); +x_7 = l_Lean_ForEachExprWhere_visit___at_Lean_Elab_Eqns_simpEqnType_collect___spec__2___rarg(x_1, x_2, x_3, x_6, x_5); +return x_7; +} +} static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_simpEqnType___spec__1___lambda__1___closed__1() { _start: { @@ -13953,7 +14213,7 @@ lean_object* x_40; lean_object* x_41; lean_object* x_42; x_40 = lean_ctor_get(x_38, 0); x_41 = lean_ctor_get(x_38, 1); lean_inc(x_2); -x_42 = l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(x_40, x_2); +x_42 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__7(x_40, x_2); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; @@ -14791,7 +15051,7 @@ lean_inc(x_176); lean_inc(x_175); lean_dec(x_38); lean_inc(x_2); -x_177 = l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(x_175, x_2); +x_177 = l_Lean_HashMapImp_find_x3f___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__7(x_175, x_2); if (lean_obj_tag(x_177) == 0) { lean_object* x_178; @@ -15683,7 +15943,7 @@ x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); lean_inc(x_17); -x_21 = l_Lean_HashMap_insert___at_Lean_ForEachExpr_visit___spec__3(x_19, x_2, x_17); +x_21 = l_Lean_HashMap_insert___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__1(x_19, x_2, x_17); x_22 = lean_st_ref_set(x_3, x_21, x_20); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) @@ -15719,7 +15979,7 @@ x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); lean_inc(x_27); -x_31 = l_Lean_HashMap_insert___at_Lean_ForEachExpr_visit___spec__3(x_29, x_2, x_27); +x_31 = l_Lean_HashMap_insert___at_Lean_Elab_Term_ContainsPendingMVar_visit___spec__1(x_29, x_2, x_27); x_32 = lean_st_ref_set(x_3, x_31, x_30); x_33 = lean_ctor_get(x_32, 1); lean_inc(x_33); @@ -16151,79 +16411,79 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_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_EXPORT lean_object* l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__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: { -if (lean_obj_tag(x_2) == 0) +if (lean_obj_tag(x_3) == 0) { -lean_object* x_10; lean_object* x_11; +lean_object* x_11; lean_object* x_12; +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -return x_11; +lean_dec(x_2); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_2, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_2, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_3, 0); lean_inc(x_13); -lean_dec(x_2); +x_14 = lean_ctor_get(x_3, 1); +lean_inc(x_14); +lean_dec(x_3); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_1); -x_14 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_14) == 0) +lean_inc(x_2); +x_15 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_2 = x_13; -x_9 = x_15; +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_3 = x_14; +x_10 = x_16; goto _start; } else { -uint8_t x_17; -lean_dec(x_13); +uint8_t x_18; +lean_dec(x_14); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_14); -if (x_17 == 0) +lean_dec(x_2); +x_18 = !lean_is_exclusive(x_15); +if (x_18 == 0) { -return x_14; +return x_15; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_14, 0); -x_19 = lean_ctor_get(x_14, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_14); -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_15); +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; } } } @@ -16381,572 +16641,582 @@ x_10 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn(x_1, x return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkEqnTypes_go___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) { _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); -lean_inc(x_6); lean_inc(x_2); lean_inc(x_1); -x_11 = l_Lean_Elab_Eqns_splitMatch_x3f(x_1, x_2, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); +x_12 = l_Lean_Elab_Eqns_splitMatch_x3f(x_1, x_2, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; -lean_dec(x_4); -lean_dec(x_2); -x_13 = lean_ctor_get(x_11, 1); +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn(x_1, x_5, x_6, x_7, x_8, x_9, x_13); -return x_14; +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_5); +lean_dec(x_2); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn(x_1, x_6, x_7, x_8, x_9, x_10, x_14); +return x_15; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_dec(x_1); -x_15 = lean_ctor_get(x_11, 1); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_ctor_get(x_12, 0); +x_16 = lean_ctor_get(x_12, 1); lean_inc(x_16); lean_dec(x_12); -x_17 = l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__2(x_2, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_15); -if (lean_obj_tag(x_17) == 0) +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec(x_13); +x_18 = l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__2(x_3, x_2, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_18) == 0) { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -return x_17; +return x_18; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +lean_inc(x_21); lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_dec(x_18); +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; -x_22 = !lean_is_exclusive(x_17); -if (x_22 == 0) +uint8_t x_23; +x_23 = !lean_is_exclusive(x_18); +if (x_23 == 0) { -return x_17; +return x_18; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_17, 0); -x_24 = lean_ctor_get(x_17, 1); +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_18, 0); +x_25 = lean_ctor_get(x_18, 1); +lean_inc(x_25); lean_inc(x_24); -lean_inc(x_23); -lean_dec(x_17); -x_25 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -return x_25; +lean_dec(x_18); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +return x_26; } } } } else { -uint8_t x_26; +uint8_t x_27; +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_26 = !lean_is_exclusive(x_11); -if (x_26 == 0) +x_27 = !lean_is_exclusive(x_12); +if (x_27 == 0) { -return x_11; +return x_12; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_11, 0); -x_28 = lean_ctor_get(x_11, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_12, 0); +x_29 = lean_ctor_get(x_12, 1); +lean_inc(x_29); lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_11); -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_dec(x_12); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -lean_dec(x_3); +lean_object* x_12; +lean_dec(x_4); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_1); -x_11 = l_Lean_MVarId_getType_x27(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_12 = l_Lean_MVarId_getType_x27(x_1, 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___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch(x_12, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) +x_15 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_shouldUseSimpMatch(x_13, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; uint8_t x_16; -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_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_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_box(0); -x_19 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2(x_1, x_2, x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -return x_19; +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_Eqns_mkEqnTypes_go___lambda__2(x_1, x_2, x_3, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +return x_20; } else { -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_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_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_1); -x_21 = l_Lean_Elab_Eqns_simpMatch_x3f(x_1, x_6, x_7, x_8, x_9, x_20); -if (lean_obj_tag(x_21) == 0) -{ -lean_object* x_22; -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); +x_22 = l_Lean_Elab_Eqns_simpMatch_x3f(x_1, x_7, x_8, x_9, x_10, x_21); if (lean_obj_tag(x_22) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_21, 1); +lean_object* x_23; +x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_box(0); -x_25 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2(x_1, x_2, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_23); -return x_25; +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = lean_box(0); +x_26 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2(x_1, x_2, x_3, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_26; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_dec(x_1); -x_26 = lean_ctor_get(x_21, 1); -lean_inc(x_26); -lean_dec(x_21); -x_27 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); lean_inc(x_27); lean_dec(x_22); -x_28 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_2, x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -if (lean_obj_tag(x_28) == 0) +x_28 = lean_ctor_get(x_23, 0); +lean_inc(x_28); +lean_dec(x_23); +x_29 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_3, x_2, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_27); +if (lean_obj_tag(x_29) == 0) { -uint8_t x_29; -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) { -return x_28; +return x_29; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_28, 0); -x_31 = lean_ctor_get(x_28, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_28); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_29); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -uint8_t x_33; -x_33 = !lean_is_exclusive(x_28); -if (x_33 == 0) +uint8_t x_34; +x_34 = !lean_is_exclusive(x_29); +if (x_34 == 0) { -return x_28; +return x_29; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_28, 0); -x_35 = lean_ctor_get(x_28, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_29, 0); +x_36 = lean_ctor_get(x_29, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_28); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_29); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } } else { -uint8_t x_37; +uint8_t x_38; +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_37 = !lean_is_exclusive(x_21); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_22); +if (x_38 == 0) { -return x_21; +return x_22; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_21, 0); -x_39 = lean_ctor_get(x_21, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_22, 0); +x_40 = lean_ctor_get(x_22, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_21); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_22); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } } else { -uint8_t x_41; +uint8_t x_42; +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_41 = !lean_is_exclusive(x_14); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_15); +if (x_42 == 0) { -return x_14; +return x_15; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_14, 0); -x_43 = lean_ctor_get(x_14, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_15, 0); +x_44 = lean_ctor_get(x_15, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_14); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec(x_15); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_45; +uint8_t 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); -lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_45 = !lean_is_exclusive(x_11); -if (x_45 == 0) +x_46 = !lean_is_exclusive(x_12); +if (x_46 == 0) { -return x_11; +return x_12; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_11, 0); -x_47 = lean_ctor_get(x_11, 1); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_12, 0); +x_48 = lean_ctor_get(x_12, 1); +lean_inc(x_48); lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_11); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +lean_dec(x_12); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkEqnTypes_go___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, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_11; -lean_dec(x_3); +lean_object* x_12; +lean_dec(x_4); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -lean_inc(x_6); lean_inc(x_1); -x_11 = l_Lean_Elab_Eqns_expandRHS_x3f(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); +x_12 = l_Lean_Elab_Eqns_expandRHS_x3f(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; -x_13 = lean_ctor_get(x_11, 1); +lean_object* x_13; +x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_box(0); -x_15 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(x_1, x_2, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_15; +if (lean_obj_tag(x_13) == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_box(0); +x_16 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(x_1, x_2, x_3, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_16; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_dec(x_1); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_ctor_get(x_12, 0); +x_17 = lean_ctor_get(x_12, 1); lean_inc(x_17); lean_dec(x_12); -x_18 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_2, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_16); -if (lean_obj_tag(x_18) == 0) +x_18 = lean_ctor_get(x_13, 0); +lean_inc(x_18); +lean_dec(x_13); +x_19 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_3, x_2, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +if (lean_obj_tag(x_19) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -return x_18; +return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_18, 0); -x_21 = lean_ctor_get(x_18, 1); +lean_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_inc(x_20); -lean_dec(x_18); -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; +lean_dec(x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } else { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_18); -if (x_23 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -return x_18; +return x_19; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_18, 0); -x_25 = lean_ctor_get(x_18, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_18); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } } else { -uint8_t x_27; +uint8_t x_28; +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_27 = !lean_is_exclusive(x_11); -if (x_27 == 0) +x_28 = !lean_is_exclusive(x_12); +if (x_28 == 0) { -return x_11; +return x_12; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_11, 0); -x_29 = lean_ctor_get(x_11, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_12, 0); +x_30 = lean_ctor_get(x_12, 1); +lean_inc(x_30); lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_11); -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_dec(x_12); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +return x_31; } } } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(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) { _start: { -lean_object* x_11; -lean_dec(x_3); +lean_dec(x_4); +if (x_3 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__4(x_1, x_2, x_3, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; +} +else +{ +lean_object* x_14; +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 = l_Lean_Elab_Eqns_tryURefl(x_1, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +x_14 = l_Lean_Elab_Eqns_tryURefl(x_1, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) +lean_object* x_15; uint8_t x_16; +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_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_box(0); -x_16 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__4(x_1, x_2, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -return x_16; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_box(0); +x_19 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__4(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_19; } else { -lean_object* x_17; lean_object* x_18; -lean_dec(x_4); +lean_object* x_20; lean_object* x_21; +lean_dec(x_5); lean_dec(x_2); -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); -x_18 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn(x_1, x_5, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_18) == 0) +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +x_21 = l___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn(x_1, x_6, x_7, x_8, x_9, x_10, x_20); +if (lean_obj_tag(x_21) == 0) { -uint8_t x_19; -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 0); +lean_dec(x_23); +x_24 = lean_box(0); +lean_ctor_set(x_21, 0, x_24); +return x_21; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_21, 1); +lean_inc(x_25); +lean_dec(x_21); +x_26 = lean_box(0); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +return x_27; } } else { -uint8_t x_25; -x_25 = !lean_is_exclusive(x_18); -if (x_25 == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) { -return x_18; +return x_21; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_18, 0); -x_27 = lean_ctor_get(x_18, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_18); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -return x_28; +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); +lean_inc(x_29); +lean_dec(x_21); +x_31 = lean_alloc_ctor(1, 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_29; +uint8_t x_32; +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_29 = !lean_is_exclusive(x_11); -if (x_29 == 0) +x_32 = !lean_is_exclusive(x_14); +if (x_32 == 0) { -return x_11; +return x_14; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_11, 0); -x_31 = lean_ctor_get(x_11, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_11); -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_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_14, 0); +x_34 = lean_ctor_get(x_14, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_14); +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; +} } } } @@ -16987,51 +17257,51 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go(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_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = l_Lean_Elab_Eqns_mkEqnTypes_go___closed__2; -x_11 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__1(x_10, 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_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = l_Lean_Elab_Eqns_mkEqnTypes_go___closed__2; +x_12 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__1(x_11, 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_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = lean_box(0); -x_16 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(x_2, x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -return x_16; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +x_16 = lean_box(0); +x_17 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(x_3, x_2, x_1, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_15); +return x_17; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_17 = lean_ctor_get(x_11, 1); -lean_inc(x_17); -lean_dec(x_11); -lean_inc(x_2); -x_18 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_18, 0, x_2); -x_19 = l_Lean_Elab_Eqns_mkEqnTypes_go___closed__4; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__9; -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_Eqns_mkEqnTypes_go___spec__3(x_10, x_22, 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_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; 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_18 = lean_ctor_get(x_12, 1); +lean_inc(x_18); +lean_dec(x_12); +lean_inc(x_3); +x_19 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_19, 0, x_3); +x_20 = l_Lean_Elab_Eqns_mkEqnTypes_go___closed__4; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_21, 1, x_19); +x_22 = l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__9; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +x_24 = l_Lean_addTrace___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__3(x_11, x_23, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(x_2, x_1, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_25); -return x_26; +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(x_3, x_2, x_1, x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +return x_27; } } } @@ -17049,6 +17319,16 @@ lean_dec(x_2); return x_9; } } +LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_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) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_List_forM___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__2(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Eqns_mkEqnTypes_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { @@ -17060,93 +17340,145 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_10; +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__1(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); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_3); +lean_dec(x_3); +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; lean_object* x_13; +x_12 = lean_unbox(x_3); +lean_dec(x_3); +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__3(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_10; -x_10 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__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_3); lean_dec(x_3); -lean_dec(x_2); -return x_10; +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__4(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes_go___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_Eqns_mkEqnTypes_go___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_11; -x_11 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__2(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_3); lean_dec(x_3); -return x_11; +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go___lambda__5(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes(lean_object* x_1, lean_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_Eqns_mkEqnTypes_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) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_8 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___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); -lean_inc(x_10); -lean_inc(x_1); -x_12 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_1, x_2, x_1, x_10, x_3, x_4, x_5, x_6, x_11); -if (lean_obj_tag(x_12) == 0) +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes(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_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_10, x_13); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___closed__1; +x_10 = lean_st_mk_ref(x_9, x_8); +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_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +lean_inc(x_11); +lean_inc(x_2); +x_13 = l_Lean_Elab_Eqns_mkEqnTypes_go(x_1, x_2, x_3, x_2, x_11, x_4, x_5, x_6, x_7, x_12); +if (lean_obj_tag(x_13) == 0) { -return x_14; +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_11, x_14); +lean_dec(x_11); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_14, 0); -x_17 = lean_ctor_get(x_14, 1); +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_inc(x_16); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_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_19; -lean_dec(x_10); -x_19 = !lean_is_exclusive(x_12); -if (x_19 == 0) +uint8_t x_20; +lean_dec(x_11); +x_20 = !lean_is_exclusive(x_13); +if (x_20 == 0) { -return x_12; +return x_13; } else { -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_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_13, 0); +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); lean_inc(x_21); -lean_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; +lean_dec(x_13); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} } } } +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkEqnTypes___boxed(lean_object* x_1, lean_object* x_2, 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_Elab_Eqns_mkEqnTypes(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_10; +} } LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, size_t x_4, size_t x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: @@ -21564,26 +21896,8 @@ return x_10; static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__1() { _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_Elab_Eqns_removeUnusedEqnHypotheses___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__3() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__2; +x_1 = l_Lean_Elab_Eqns_tryURefl___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); @@ -21591,7 +21905,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4() { +static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__2() { _start: { lean_object* x_1; lean_object* x_2; @@ -21600,23 +21914,23 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5() { +static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4; +x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6() { +static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4() { _start: { size_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 5; -x_2 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5; -x_3 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4; +x_2 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__3; +x_3 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__2; 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); @@ -21627,12 +21941,12 @@ lean_ctor_set_usize(x_5, 4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7() { +static lean_object* _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__3; -x_2 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6; +x_1 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__1; +x_2 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -21644,7 +21958,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses(lean_object* { lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___closed__1; -x_7 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7; +x_7 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5; lean_inc(x_2); lean_inc(x_1); x_8 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses_go(x_1, x_2, x_1, x_2, x_6, x_7, x_3, x_4, x_5); @@ -25001,46 +25315,42 @@ return x_67; } } } -static lean_object* _init_l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_Meta_tactic_hygienic; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___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_Eqns_mkUnfoldEq___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -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_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1; -x_11 = 0; -x_12 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_9, x_10, x_11); -lean_ctor_set(x_5, 2, x_12); -x_13 = l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(x_1, x_2, x_11, x_3, x_4, x_5, x_6, x_7); -return x_13; +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_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); +x_16 = 0; +x_17 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_3, x_4, x_16, x_5, x_6, x_8, x_9, x_10); +return x_17; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t 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_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); -x_26 = lean_ctor_get_uint8(x_5, sizeof(void*)*11 + 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 3); +x_21 = lean_ctor_get(x_8, 5); +x_22 = lean_ctor_get(x_8, 6); +x_23 = lean_ctor_get(x_8, 7); +x_24 = lean_ctor_get(x_8, 8); +x_25 = lean_ctor_get(x_8, 9); +x_26 = lean_ctor_get(x_8, 10); +x_27 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +lean_inc(x_26); +lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); @@ -25048,49 +25358,208 @@ 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_27 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1; -x_28 = 0; -x_29 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_16, x_27, x_28); +lean_dec(x_8); +x_28 = l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1; +x_29 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_28); x_30 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_30, 0, x_14); -lean_ctor_set(x_30, 1, x_15); -lean_ctor_set(x_30, 2, x_29); -lean_ctor_set(x_30, 3, x_17); -lean_ctor_set(x_30, 4, x_18); -lean_ctor_set(x_30, 5, x_19); -lean_ctor_set(x_30, 6, x_20); -lean_ctor_set(x_30, 7, x_21); -lean_ctor_set(x_30, 8, x_22); -lean_ctor_set(x_30, 9, x_23); -lean_ctor_set(x_30, 10, x_24); -lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_25); -lean_ctor_set_uint8(x_30, sizeof(void*)*11 + 1, x_26); -x_31 = l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(x_1, x_2, x_28, x_3, x_4, x_30, x_6, x_7); -return x_31; +lean_ctor_set(x_30, 0, x_18); +lean_ctor_set(x_30, 1, x_19); +lean_ctor_set(x_30, 2, x_1); +lean_ctor_set(x_30, 3, x_20); +lean_ctor_set(x_30, 4, x_29); +lean_ctor_set(x_30, 5, x_21); +lean_ctor_set(x_30, 6, x_22); +lean_ctor_set(x_30, 7, x_23); +lean_ctor_set(x_30, 8, x_24); +lean_ctor_set(x_30, 9, x_25); +lean_ctor_set(x_30, 10, x_26); +lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_27); +lean_ctor_set_uint8(x_30, sizeof(void*)*11 + 1, x_2); +x_31 = 0; +x_32 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_3, x_4, x_31, x_5, x_6, x_30, x_9, x_10); +return x_32; +} +} +} +static lean_object* _init_l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_tactic_hygienic; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +x_10 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1; +x_11 = 0; +x_12 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_9, x_10, x_11); +x_13 = l_Lean_Elab_Eqns_tryURefl___closed__2; +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_12, x_13); +x_15 = lean_st_ref_get(x_7, x_8); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_Kernel_isDiagnosticsEnabled(x_18); +lean_dec(x_18); +if (x_19 == 0) +{ +if (x_14 == 0) +{ +uint8_t x_53; +x_53 = 1; +x_20 = x_53; +goto block_52; +} +else +{ +x_20 = x_11; +goto block_52; +} +} +else +{ +if (x_14 == 0) +{ +x_20 = x_11; +goto block_52; +} +else +{ +uint8_t x_54; +x_54 = 1; +x_20 = x_54; +goto block_52; +} +} +block_52: +{ +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_st_ref_take(x_7, x_17); +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = !lean_is_exclusive(x_22); +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_25 = lean_ctor_get(x_22, 0); +x_26 = lean_ctor_get(x_22, 4); +lean_dec(x_26); +x_27 = l_Lean_Kernel_enableDiag(x_25, x_14); +x_28 = lean_unsigned_to_nat(0u); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_28); +lean_inc(x_29); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_22, 4, x_30); +lean_ctor_set(x_22, 0, x_27); +x_31 = lean_st_ref_set(x_7, x_22, x_23); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2(x_12, x_14, x_1, x_2, x_4, x_5, x_33, x_6, x_7, 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_35 = lean_ctor_get(x_22, 0); +x_36 = lean_ctor_get(x_22, 1); +x_37 = lean_ctor_get(x_22, 2); +x_38 = lean_ctor_get(x_22, 3); +x_39 = lean_ctor_get(x_22, 5); +x_40 = lean_ctor_get(x_22, 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_dec(x_22); +x_41 = l_Lean_Kernel_enableDiag(x_35, x_14); +x_42 = lean_unsigned_to_nat(0u); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_3); +lean_ctor_set(x_43, 1, x_42); +lean_inc(x_43); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_43); +x_45 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_45, 0, x_41); +lean_ctor_set(x_45, 1, x_36); +lean_ctor_set(x_45, 2, x_37); +lean_ctor_set(x_45, 3, x_38); +lean_ctor_set(x_45, 4, x_44); +lean_ctor_set(x_45, 5, x_39); +lean_ctor_set(x_45, 6, x_40); +x_46 = lean_st_ref_set(x_7, x_45, x_23); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_box(0); +x_49 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2(x_12, x_14, x_1, x_2, x_4, x_5, x_48, x_6, x_7, x_47); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; +lean_dec(x_3); +x_50 = lean_box(0); +x_51 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2(x_12, x_14, x_1, x_2, x_4, x_5, x_50, x_6, x_7, x_17); +return x_51; +} } } } LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq(lean_object* x_1, 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_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; x_8 = lean_ctor_get(x_2, 3); lean_inc(x_8); x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Eqns_mkUnfoldEq___lambda__1), 9, 2); lean_closure_set(x_9, 0, x_2); lean_closure_set(x_9, 1, x_1); -x_10 = lean_alloc_closure((void*)(l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2), 7, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7; -x_12 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___closed__1; -x_13 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_11, x_12, x_10, x_3, x_4, x_5, x_6, x_7); -return x_13; +x_10 = l_Lean_Elab_Eqns_tryURefl___closed__4; +x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3), 8, 3); +lean_closure_set(x_11, 0, x_8); +lean_closure_set(x_11, 1, x_9); +lean_closure_set(x_11, 2, x_10); +x_12 = l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5; +x_13 = l_Lean_RBTree_toArray___at___private_Lean_Elab_PreDefinition_Eqns_0__Lean_Elab_Eqns_saveEqn_collectDeps___spec__1___closed__1; +x_14 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_12, x_13, x_11, x_3, x_4, x_5, x_6, x_7); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_mkUnfoldEq___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_2); +lean_dec(x_2); +x_12 = l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_7); +return x_12; } } LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { @@ -25234,7 +25703,7 @@ return x_34; } } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1() { _start: { lean_object* x_1; @@ -25242,18 +25711,18 @@ x_1 = lean_mk_string_from_bytes("unfoldEqn", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__1; x_2 = l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__2; -x_3 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1; +x_3 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3() { _start: { lean_object* x_1; @@ -25261,27 +25730,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4; x_2 = l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6() { _start: { lean_object* x_1; @@ -25289,17 +25758,17 @@ x_1 = lean_mk_string_from_bytes("Eqns", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8() { _start: { lean_object* x_1; @@ -25307,17 +25776,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10() { _start: { lean_object* x_1; @@ -25325,37 +25794,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12; x_2 = l_Lean_Elab_Eqns_simpEqnType___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14() { _start: { lean_object* x_1; @@ -25363,27 +25832,27 @@ x_1 = lean_mk_string_from_bytes("PreDefinition", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17() { _start: { lean_object* x_1; @@ -25391,33 +25860,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16; -x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17; +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19() { +static lean_object* _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18; -x_2 = lean_unsigned_to_nat(6458u); +x_1 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18; +x_2 = lean_unsigned_to_nat(6487u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487_(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_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2; +x_2 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19; +x_4 = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -25544,8 +26013,20 @@ l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__9 = _init_l_Lean_Elab_Eqns_splitMat lean_mark_persistent(l_Lean_Elab_Eqns_splitMatch_x3f_go___closed__9); l_Lean_Elab_Eqns_splitMatch_x3f___lambda__1___closed__1 = _init_l_Lean_Elab_Eqns_splitMatch_x3f___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Eqns_splitMatch_x3f___lambda__1___closed__1); +l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1 = _init_l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___lambda__1___closed__1); l_Lean_Elab_Eqns_tryURefl___closed__1 = _init_l_Lean_Elab_Eqns_tryURefl___closed__1(); lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__1); +l_Lean_Elab_Eqns_tryURefl___closed__2 = _init_l_Lean_Elab_Eqns_tryURefl___closed__2(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__2); +l_Lean_Elab_Eqns_tryURefl___closed__3 = _init_l_Lean_Elab_Eqns_tryURefl___closed__3(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__3); +l_Lean_Elab_Eqns_tryURefl___closed__4 = _init_l_Lean_Elab_Eqns_tryURefl___closed__4(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__4); +l_Lean_Elab_Eqns_tryURefl___closed__5 = _init_l_Lean_Elab_Eqns_tryURefl___closed__5(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__5); +l_Lean_Elab_Eqns_tryURefl___closed__6 = _init_l_Lean_Elab_Eqns_tryURefl___closed__6(); +lean_mark_persistent(l_Lean_Elab_Eqns_tryURefl___closed__6); l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__1 = _init_l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__1(); lean_mark_persistent(l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__1); l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__2 = _init_l_panic___at_Lean_Elab_Eqns_simpEqnType_collect___spec__1___rarg___closed__2(); @@ -25625,10 +26106,6 @@ l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4 = _init_l_Lean_Elab_Eqns_ lean_mark_persistent(l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__4); l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5 = _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5(); lean_mark_persistent(l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__5); -l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6 = _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6(); -lean_mark_persistent(l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__6); -l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7 = _init_l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7(); -lean_mark_persistent(l_Lean_Elab_Eqns_removeUnusedEqnHypotheses___closed__7); l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__1 = _init_l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__1); l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__2 = _init_l_Lean_Elab_Eqns_deltaLHS___lambda__2___closed__2(); @@ -25669,47 +26146,47 @@ l_Lean_Elab_Eqns_mkUnfoldProof___closed__3 = _init_l_Lean_Elab_Eqns_mkUnfoldProo lean_mark_persistent(l_Lean_Elab_Eqns_mkUnfoldProof___closed__3); l_Lean_Elab_Eqns_mkUnfoldProof___closed__4 = _init_l_Lean_Elab_Eqns_mkUnfoldProof___closed__4(); lean_mark_persistent(l_Lean_Elab_Eqns_mkUnfoldProof___closed__4); -l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1 = _init_l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Elab_Eqns_mkUnfoldEq___lambda__2___closed__1); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__1); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__2); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__3); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__4); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__5); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__6); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__7); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__8); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__9); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__10); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__11); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__12); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__13); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__14); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__15); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__16); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__17); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__18); -l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19(); -lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458____closed__19); -if (builtin) {res = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6458_(lean_io_mk_world()); +l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1 = _init_l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Elab_Eqns_mkUnfoldEq___lambda__3___closed__1); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__1); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__2); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__3); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__4); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__5); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__6); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__7); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__8); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__9); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__10); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__11); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__12); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__13); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__14); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__15); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__16); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__17); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__18); +l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19 = _init_l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19(); +lean_mark_persistent(l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487____closed__19); +if (builtin) {res = l_Lean_Elab_Eqns_initFn____x40_Lean_Elab_PreDefinition_Eqns___hyg_6487_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c index 05cac5218ca7..0ce6ab57ae4e 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/BRecOn.c @@ -136,6 +136,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_BRecOn_0 lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Structural_mkBRecOn___spec__2___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__41___lambda__3___closed__4; +lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__44___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*); size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_toBelowAux___lambda__5___closed__8; @@ -147,6 +148,7 @@ extern lean_object* l_Lean_binductionOnSuffix; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__33___rarg___boxed(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_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__25___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkBRecOn___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*); +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___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_isTracingEnabledFor___at_Lean_Elab_Structural_mkBRecOn___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -169,7 +171,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_PreDefinitio LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__43___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__7(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_withBelowDict(lean_object*); -uint8_t l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__36___lambda__2___boxed(lean_object**); @@ -223,7 +224,6 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefi static lean_object* l_Lean_getConstInfo___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__41___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_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_replaceRecApps_loop___spec__23___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__5___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_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_toBelowAux___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_withBelowDict___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*); @@ -1922,7 +1922,7 @@ lean_closure_set(x_18, 1, x_1); lean_closure_set(x_18, 2, x_4); lean_closure_set(x_18, 3, x_5); x_19 = l_Lean_Expr_withAppAux___at___private_Lean_Elab_PreDefinition_Structural_BRecOn_0__Lean_Elab_Structural_withBelowDict___spec__1___rarg___lambda__3___closed__1; -x_20 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_Meta_FunInfo_0__Lean_Meta_getFunInfoAux___spec__5___rarg(x_16, x_19, x_18, x_7, x_8, x_9, x_10, x_17); +x_20 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_16, x_19, x_18, x_7, x_8, x_9, x_10, x_17); return x_20; } else @@ -12789,7 +12789,7 @@ if (x_9 == 0) { lean_object* x_12; uint8_t x_13; x_12 = lean_ctor_get(x_1, 7); -x_13 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_12, x_8); +x_13 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_12, x_8); if (x_13 == 0) { lean_object* x_14; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c index 0ba741e07de9..b799c42b2716 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/Eqns.c @@ -16,14 +16,11 @@ extern "C" { lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_instInhabitedEqnInfo___closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___closed__2; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(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_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__16; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__5; uint32_t l_UInt32_ofNatTruncate(lean_object*); static lean_object* l_Lean_Elab_Structural_registerEqnsInfo___closed__1; @@ -31,10 +28,9 @@ static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_ LEAN_EXPORT lean_object* l_Lean_Elab_Structural_getUnfoldFor_x3f___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___lambda__3___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_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9; lean_object* l_Lean_Elab_Eqns_simpMatch_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1; uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19; LEAN_EXPORT lean_object* lean_get_structural_rec_arg_pos(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_instInhabitedEqnInfo; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_registerEqnsInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -44,63 +40,84 @@ static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_ lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___closed__1; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6; extern lean_object* l_Lean_Meta_tactic_hygienic; lean_object* l_Lean_Meta_registerGetUnfoldEqnFn(lean_object*, lean_object*); lean_object* l_Lean_Elab_Eqns_tryContradiction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__4; +extern lean_object* l_Lean_maxRecDepth; 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___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15; -LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437_(lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); +static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_eqnInfoExt; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Structural_mkEqns___closed__1; lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___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*); +static lean_object* l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___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_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12; lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1; lean_object* l_Lean_Elab_Eqns_deltaRHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3; lean_object* l_Lean_Meta_mkEq(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_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16; extern lean_object* l_Lean_instInhabitedExpr; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10; +lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___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*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__1; lean_object* l_Lean_Name_str___override(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___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__5; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__7; +extern lean_object* l_Lean_diagnostics; +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195_(lean_object*); +static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20; lean_object* l_Lean_Elab_Eqns_simpIf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Elab_Eqns_mkEqnTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4; +lean_object* l_Lean_Elab_Eqns_mkEqnTypes(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_splitTarget_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instBEqOrigin___boxed(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_getEqnsFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_eqnThmSuffixBase; lean_object* lean_name_append_index_after(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__14; static lean_object* l_Lean_Elab_Structural_instInhabitedEqnInfo___closed__1; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_registerGetEqnsFn(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -109,61 +126,56 @@ lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Structural_mkEqns___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10; -lean_object* l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_mkEqns___closed__2; lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_ensureEqnReservedNamesAvailable(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3; -lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__12; lean_object* l_Lean_Elab_Eqns_whnfReducibleLHS_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_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go(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_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12; lean_object* l_Lean_Meta_instHashableOrigin___boxed(lean_object*); extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14; lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Meta_casesOnStuckLHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1; lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_simpTargetStar(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_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_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194_(lean_object*); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5; -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13; lean_object* l_Lean_Elab_Eqns_tryURefl(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*); 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_List_forM___at___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___lambda__2(lean_object*, uint8_t, 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_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8; +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__11; lean_object* l_Lean_Elab_Eqns_deltaLHS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static uint32_t l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__2; lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_getUnfoldFor_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Structural_registerEqnsInfo___closed__2; lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__3; static lean_object* l_Lean_Elab_Structural_instInhabitedEqnInfo___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__13; +static lean_object* l_Lean_Elab_Structural_mkEqns___closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__7; static lean_object* l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17; static lean_object* _init_l_Lean_Elab_Structural_instInhabitedEqnInfo___closed__1() { @@ -462,21 +474,23 @@ return x_4; static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__9() { _start: { -lean_object* x_1; uint32_t x_2; lean_object* x_3; uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_1; uint32_t x_2; lean_object* x_3; uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_1 = lean_box(0); x_2 = 0; x_3 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__1; x_4 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__2; x_5 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; x_6 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__8; -x_7 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_5); -lean_ctor_set(x_7, 2, x_6); -lean_ctor_set(x_7, 3, x_1); -lean_ctor_set_uint32(x_7, sizeof(void*)*4, x_4); -lean_ctor_set_uint32(x_7, sizeof(void*)*4 + 4, x_2); -return x_7; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_8, 0, x_3); +lean_ctor_set(x_8, 1, x_5); +lean_ctor_set(x_8, 2, x_6); +lean_ctor_set(x_8, 3, x_1); +lean_ctor_set(x_8, 4, x_7); +lean_ctor_set_uint32(x_8, sizeof(void*)*5, x_4); +lean_ctor_set_uint32(x_8, sizeof(void*)*5 + 4, x_2); +return x_8; } } static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__10() { @@ -510,21 +524,46 @@ return x_3; static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__13() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__12; +x_2 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__7; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__12; +x_2 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__13; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("failed to generate equational theorem for '", 43); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__14() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__13; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17() { _start: { lean_object* x_1; @@ -532,16 +571,16 @@ x_1 = lean_mk_string_from_bytes("'\n", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__16() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__15; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19() { _start: { lean_object* x_1; @@ -549,11 +588,11 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -655,7 +694,7 @@ lean_dec(x_23); x_26 = lean_box(0); x_27 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__9; x_28 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; -x_29 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__12; +x_29 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__14; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); @@ -763,11 +802,11 @@ x_49 = lean_ctor_get(x_47, 1); lean_inc(x_49); lean_dec(x_47); x_50 = l_Lean_MessageData_ofName(x_2); -x_51 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__14; +x_51 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__16; 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___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__16; +x_53 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18; x_54 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_54, 0, x_52); lean_ctor_set(x_54, 1, x_53); @@ -776,7 +815,7 @@ lean_ctor_set(x_55, 0, x_1); 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___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18; +x_57 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20; x_58 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_58, 0, x_56); lean_ctor_set(x_58, 1, x_57); @@ -1420,7 +1459,7 @@ x_17 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structur x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18; +x_19 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20; x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); @@ -1730,7 +1769,7 @@ x_17 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structur x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_17); lean_ctor_set(x_18, 1, x_16); -x_19 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18; +x_19 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20; x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); @@ -2064,7 +2103,7 @@ 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_inc(x_23); x_43 = l_Lean_MessageData_ofExpr(x_23); -x_44 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18; +x_44 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20; x_45 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_45, 0, x_44); lean_ctor_set(x_45, 1, x_43); @@ -2244,7 +2283,7 @@ lean_inc(x_4); x_15 = l_Lean_Meta_mkEq(x_14, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); @@ -2261,190 +2300,176 @@ lean_dec(x_19); x_22 = l_Lean_Elab_Structural_mkEqns___lambda__1___closed__1; x_23 = lean_array_push(x_22, x_12); x_24 = l_Lean_Expr_mvarId_x21(x_20); -x_25 = l_Lean_Elab_Eqns_mkEqnTypes(x_23, x_24, x_4, x_5, x_6, x_7, x_21); -return x_25; +x_25 = 1; +x_26 = l_Lean_Elab_Eqns_mkEqnTypes(x_25, x_23, x_24, x_4, x_5, x_6, x_7, x_21); +return x_26; } else { -uint8_t x_26; +uint8_t x_27; lean_dec(x_12); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_26 = !lean_is_exclusive(x_15); -if (x_26 == 0) +x_27 = !lean_is_exclusive(x_15); +if (x_27 == 0) { return x_15; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_15, 0); -x_28 = lean_ctor_get(x_15, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_15, 0); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); lean_inc(x_28); -lean_inc(x_27); lean_dec(x_15); -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; +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } } } -static lean_object* _init_l_Lean_Elab_Structural_mkEqns___closed__1() { +static lean_object* _init_l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_Meta_tactic_hygienic; +x_1 = l_Lean_maxRecDepth; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns(lean_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_Structural_mkEqns___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -lean_dec(x_1); -x_8 = lean_ctor_get(x_7, 3); -lean_inc(x_8); -lean_inc(x_7); -x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Structural_mkEqns___lambda__1), 8, 1); -lean_closure_set(x_9, 0, x_7); -x_10 = 1; -x_11 = lean_box(x_10); -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed), 8, 3); -lean_closure_set(x_12, 0, x_8); -lean_closure_set(x_12, 1, x_9); -lean_closure_set(x_12, 2, x_11); -x_13 = !lean_is_exclusive(x_4); -if (x_13 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_4, 2); -x_15 = l_Lean_Elab_Structural_mkEqns___closed__1; +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_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); x_16 = 0; -x_17 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_14, x_15, x_16); -lean_ctor_set(x_4, 2, x_17); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -lean_inc(x_2); -x_18 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_12, x_16, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_18) == 0) +x_17 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_3, x_16, x_4, x_5, x_8, x_9, x_10); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_19 = lean_ctor_get(x_18, 0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_dec(x_17); +x_20 = lean_ctor_get(x_6, 0); lean_inc(x_20); +x_21 = lean_box(0); +x_22 = lean_array_get_size(x_18); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; +lean_inc(x_22); +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1(x_6, x_18, x_20, x_21, x_22, x_22, x_23, x_22, x_24, x_25, x_4, x_5, x_8, x_9, x_19); +lean_dec(x_22); lean_dec(x_18); -x_21 = lean_ctor_get(x_7, 0); -lean_inc(x_21); -x_22 = lean_box(0); -x_23 = lean_array_get_size(x_19); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_unsigned_to_nat(1u); -x_26 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; -lean_inc(x_23); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1(x_7, x_19, x_21, x_22, x_23, x_23, x_24, x_23, x_25, x_26, x_2, x_3, x_4, x_5, x_20); -lean_dec(x_23); -lean_dec(x_19); -lean_dec(x_7); -if (lean_obj_tag(x_27) == 0) +lean_dec(x_6); +if (lean_obj_tag(x_26) == 0) { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -return x_27; +return x_26; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); -lean_dec(x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +uint8_t x_31; +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) { -return x_27; +return x_26; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_inc(x_32); +lean_dec(x_26); +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_36; -lean_dec(x_4); -lean_dec(x_7); +uint8_t x_35; +lean_dec(x_8); +lean_dec(x_9); +lean_dec(x_6); lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_36 = !lean_is_exclusive(x_18); -if (x_36 == 0) +lean_dec(x_4); +x_35 = !lean_is_exclusive(x_17); +if (x_35 == 0) { -return x_18; +return x_17; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_18, 0); -x_38 = lean_ctor_get(x_18, 1); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_17, 0); +x_37 = lean_ctor_get(x_17, 1); lean_inc(x_37); -lean_dec(x_18); -x_39 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_39, 0, x_37); -lean_ctor_set(x_39, 1, x_38); -return x_39; +lean_inc(x_36); +lean_dec(x_17); +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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_40 = lean_ctor_get(x_4, 0); -x_41 = lean_ctor_get(x_4, 1); -x_42 = lean_ctor_get(x_4, 2); -x_43 = lean_ctor_get(x_4, 3); -x_44 = lean_ctor_get(x_4, 4); -x_45 = lean_ctor_get(x_4, 5); -x_46 = lean_ctor_get(x_4, 6); -x_47 = lean_ctor_get(x_4, 7); -x_48 = lean_ctor_get(x_4, 8); -x_49 = lean_ctor_get(x_4, 9); -x_50 = lean_ctor_get(x_4, 10); -x_51 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); -x_52 = lean_ctor_get_uint8(x_4, sizeof(void*)*11 + 1); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_39 = lean_ctor_get(x_8, 0); +x_40 = lean_ctor_get(x_8, 1); +x_41 = lean_ctor_get(x_8, 3); +x_42 = lean_ctor_get(x_8, 5); +x_43 = lean_ctor_get(x_8, 6); +x_44 = lean_ctor_get(x_8, 7); +x_45 = lean_ctor_get(x_8, 8); +x_46 = lean_ctor_get(x_8, 9); +x_47 = lean_ctor_get(x_8, 10); +x_48 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); @@ -2453,66 +2478,91 @@ lean_inc(x_43); lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); -lean_dec(x_4); -x_53 = l_Lean_Elab_Structural_mkEqns___closed__1; -x_54 = 0; -x_55 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_42, x_53, x_54); -x_56 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_56, 0, x_40); -lean_ctor_set(x_56, 1, x_41); -lean_ctor_set(x_56, 2, x_55); -lean_ctor_set(x_56, 3, x_43); -lean_ctor_set(x_56, 4, x_44); -lean_ctor_set(x_56, 5, x_45); -lean_ctor_set(x_56, 6, x_46); -lean_ctor_set(x_56, 7, x_47); -lean_ctor_set(x_56, 8, x_48); -lean_ctor_set(x_56, 9, x_49); -lean_ctor_set(x_56, 10, x_50); -lean_ctor_set_uint8(x_56, sizeof(void*)*11, x_51); -lean_ctor_set_uint8(x_56, sizeof(void*)*11 + 1, x_52); +lean_inc(x_39); +lean_dec(x_8); +x_49 = l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1; +x_50 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_49); +x_51 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_51, 0, x_39); +lean_ctor_set(x_51, 1, x_40); +lean_ctor_set(x_51, 2, x_1); +lean_ctor_set(x_51, 3, x_41); +lean_ctor_set(x_51, 4, x_50); +lean_ctor_set(x_51, 5, x_42); +lean_ctor_set(x_51, 6, x_43); +lean_ctor_set(x_51, 7, x_44); +lean_ctor_set(x_51, 8, x_45); +lean_ctor_set(x_51, 9, x_46); +lean_ctor_set(x_51, 10, x_47); +lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_2); +x_52 = 0; +lean_inc(x_9); +lean_inc(x_51); lean_inc(x_5); -lean_inc(x_56); -lean_inc(x_3); -lean_inc(x_2); -x_57 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_12, x_54, x_2, x_3, x_56, x_5, x_6); -if (lean_obj_tag(x_57) == 0) +lean_inc(x_4); +x_53 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_3, x_52, x_4, x_5, x_51, x_9, x_10); +if (lean_obj_tag(x_53) == 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_57, 0); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_ctor_get(x_6, 0); +lean_inc(x_56); +x_57 = lean_box(0); +x_58 = lean_array_get_size(x_54); +x_59 = lean_unsigned_to_nat(0u); +x_60 = lean_unsigned_to_nat(1u); +x_61 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_ctor_get(x_7, 0); -lean_inc(x_60); -x_61 = lean_box(0); -x_62 = lean_array_get_size(x_58); -x_63 = lean_unsigned_to_nat(0u); -x_64 = lean_unsigned_to_nat(1u); -x_65 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__3; -lean_inc(x_62); -x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1(x_7, x_58, x_60, x_61, x_62, x_62, x_63, x_62, x_64, x_65, x_2, x_3, x_56, x_5, x_59); -lean_dec(x_62); +x_62 = l_Std_Range_forIn_loop___at_Lean_Elab_Structural_mkEqns___spec__1(x_6, x_54, x_56, x_57, x_58, x_58, x_59, x_58, x_60, x_61, x_4, x_5, x_51, x_9, x_55); lean_dec(x_58); -lean_dec(x_7); -if (lean_obj_tag(x_66) == 0) +lean_dec(x_54); +lean_dec(x_6); +if (lean_obj_tag(x_62) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +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; +} else { + lean_dec_ref(x_62); + x_65 = lean_box(0); +} +if (lean_is_scalar(x_65)) { + x_66 = lean_alloc_ctor(0, 2, 0); +} else { + x_66 = x_65; +} +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; -x_67 = lean_ctor_get(x_66, 0); +x_67 = lean_ctor_get(x_62, 0); lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +x_68 = lean_ctor_get(x_62, 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; +if (lean_is_exclusive(x_62)) { + lean_ctor_release(x_62, 0); + lean_ctor_release(x_62, 1); + x_69 = x_62; } else { - lean_dec_ref(x_66); + lean_dec_ref(x_62); x_69 = lean_box(0); } if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(0, 2, 0); + x_70 = lean_alloc_ctor(1, 2, 0); } else { x_70 = x_69; } @@ -2520,19 +2570,25 @@ lean_ctor_set(x_70, 0, x_67); lean_ctor_set(x_70, 1, x_68); return x_70; } +} else { lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_71 = lean_ctor_get(x_66, 0); +lean_dec(x_51); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_71 = lean_ctor_get(x_53, 0); lean_inc(x_71); -x_72 = lean_ctor_get(x_66, 1); +x_72 = lean_ctor_get(x_53, 1); lean_inc(x_72); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_73 = x_66; +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_73 = x_53; } else { - lean_dec_ref(x_66); + lean_dec_ref(x_53); x_73 = lean_box(0); } if (lean_is_scalar(x_73)) { @@ -2545,34 +2601,168 @@ lean_ctor_set(x_74, 1, x_72); return x_74; } } +} +} +static lean_object* _init_l_Lean_Elab_Structural_mkEqns___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_tactic_hygienic; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Structural_mkEqns___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Structural_mkEqns___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__7; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; uint8_t x_24; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +lean_dec(x_1); +x_8 = lean_ctor_get(x_7, 3); +lean_inc(x_8); +lean_inc(x_7); +x_9 = lean_alloc_closure((void*)(l_Lean_Elab_Structural_mkEqns___lambda__1), 8, 1); +lean_closure_set(x_9, 0, x_7); +x_10 = 1; +x_11 = lean_box(x_10); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed), 8, 3); +lean_closure_set(x_12, 0, x_8); +lean_closure_set(x_12, 1, x_9); +lean_closure_set(x_12, 2, x_11); +x_13 = lean_ctor_get(x_4, 2); +lean_inc(x_13); +x_14 = l_Lean_Elab_Structural_mkEqns___closed__1; +x_15 = 0; +x_16 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_13, x_14, x_15); +x_17 = l_Lean_Elab_Structural_mkEqns___closed__2; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_16, x_17); +x_19 = lean_st_ref_get(x_5, x_6); +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_Kernel_isDiagnosticsEnabled(x_22); +lean_dec(x_22); +if (x_23 == 0) +{ +if (x_18 == 0) +{ +x_24 = x_10; +goto block_52; +} else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_56); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_3); -lean_dec(x_2); -x_75 = lean_ctor_get(x_57, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_57, 1); -lean_inc(x_76); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_77 = x_57; -} else { - lean_dec_ref(x_57); - x_77 = lean_box(0); +x_24 = x_15; +goto block_52; } -if (lean_is_scalar(x_77)) { - x_78 = lean_alloc_ctor(1, 2, 0); -} else { - x_78 = x_77; } -lean_ctor_set(x_78, 0, x_75); -lean_ctor_set(x_78, 1, x_76); -return x_78; +else +{ +if (x_18 == 0) +{ +x_24 = x_15; +goto block_52; +} +else +{ +x_24 = x_10; +goto block_52; +} +} +block_52: +{ +if (x_24 == 0) +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_st_ref_take(x_5, x_21); +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_26); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = lean_ctor_get(x_26, 0); +x_30 = lean_ctor_get(x_26, 4); +lean_dec(x_30); +x_31 = l_Lean_Kernel_enableDiag(x_29, x_18); +x_32 = l_Lean_Elab_Structural_mkEqns___closed__3; +lean_ctor_set(x_26, 4, x_32); +lean_ctor_set(x_26, 0, x_31); +x_33 = lean_st_ref_set(x_5, x_26, x_27); +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); +x_35 = lean_box(0); +x_36 = l_Lean_Elab_Structural_mkEqns___lambda__2(x_16, x_18, x_12, x_2, x_3, x_7, x_35, x_4, x_5, x_34); +return x_36; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_37 = lean_ctor_get(x_26, 0); +x_38 = lean_ctor_get(x_26, 1); +x_39 = lean_ctor_get(x_26, 2); +x_40 = lean_ctor_get(x_26, 3); +x_41 = lean_ctor_get(x_26, 5); +x_42 = lean_ctor_get(x_26, 6); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_26); +x_43 = l_Lean_Kernel_enableDiag(x_37, x_18); +x_44 = l_Lean_Elab_Structural_mkEqns___closed__3; +x_45 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_38); +lean_ctor_set(x_45, 2, x_39); +lean_ctor_set(x_45, 3, x_40); +lean_ctor_set(x_45, 4, x_44); +lean_ctor_set(x_45, 5, x_41); +lean_ctor_set(x_45, 6, x_42); +x_46 = lean_st_ref_set(x_5, x_45, x_27); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_box(0); +x_49 = l_Lean_Elab_Structural_mkEqns___lambda__2(x_16, x_18, x_12, x_2, x_3, x_7, x_48, x_4, x_5, x_47); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_box(0); +x_51 = l_Lean_Elab_Structural_mkEqns___lambda__2(x_16, x_18, x_12, x_2, x_3, x_7, x_50, x_4, x_5, x_21); +return x_51; } } } @@ -2601,7 +2791,18 @@ lean_dec(x_1); return x_16; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkEqns___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_2); +lean_dec(x_2); +x_12 = l_Lean_Elab_Structural_mkEqns___lambda__2(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_7); +return x_12; +} +} +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1() { _start: { lean_object* x_1; @@ -2609,7 +2810,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2() { _start: { lean_object* x_1; @@ -2617,7 +2818,7 @@ x_1 = lean_mk_string_from_bytes("Structural", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3() { _start: { lean_object* x_1; @@ -2625,23 +2826,23 @@ x_1 = lean_mk_string_from_bytes("eqnInfoExt", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____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_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1; x_2 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1; -x_3 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2; -x_4 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3; +x_3 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2; +x_4 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__4; x_3 = l_Lean_mkMapDeclarationExtension___rarg(x_2, x_1); return x_3; } @@ -2654,17 +2855,6 @@ x_1 = l_Lean_Elab_Structural_eqnInfoExt; return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_registerEqnsInfo___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__7; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Structural_registerEqnsInfo(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -2711,7 +2901,7 @@ lean_ctor_set(x_19, 0, x_18); lean_ctor_set(x_19, 1, x_2); x_20 = l_Lean_Elab_Structural_registerEqnsInfo___closed__1; x_21 = l_Lean_MapDeclarationExtension_insert___rarg(x_20, x_13, x_6, x_19); -x_22 = l_Lean_Elab_Structural_registerEqnsInfo___closed__2; +x_22 = l_Lean_Elab_Structural_mkEqns___closed__3; lean_ctor_set(x_10, 4, x_22); lean_ctor_set(x_10, 0, x_21); x_23 = lean_st_ref_set(x_4, x_10, x_11); @@ -2773,7 +2963,7 @@ lean_ctor_set(x_40, 0, x_39); lean_ctor_set(x_40, 1, x_2); x_41 = l_Lean_Elab_Structural_registerEqnsInfo___closed__1; x_42 = l_Lean_MapDeclarationExtension_insert___rarg(x_41, x_30, x_6, x_40); -x_43 = l_Lean_Elab_Structural_registerEqnsInfo___closed__2; +x_43 = l_Lean_Elab_Structural_mkEqns___closed__3; x_44 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_31); @@ -3262,7 +3452,7 @@ return x_31; } } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1() { _start: { lean_object* x_1; @@ -3270,7 +3460,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Structural_getEqnsFor_x3f), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2() { _start: { lean_object* x_1; @@ -3278,37 +3468,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Structural_getUnfoldFor_x3f), 6, 0) return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3; x_2 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6() { _start: { lean_object* x_1; @@ -3316,17 +3506,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8() { _start: { lean_object* x_1; @@ -3334,37 +3524,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10; x_2 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12() { _start: { lean_object* x_1; @@ -3372,27 +3562,27 @@ x_1 = lean_mk_string_from_bytes("PreDefinition", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15() { _start: { lean_object* x_1; @@ -3400,17 +3590,17 @@ x_1 = lean_mk_string_from_bytes("Eqns", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17() { _start: { lean_object* x_1; @@ -3418,31 +3608,31 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17; +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19() { +static lean_object* _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18; -x_2 = lean_unsigned_to_nat(1437u); +x_1 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18; +x_2 = lean_unsigned_to_nat(1438u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1; +x_2 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1; x_3 = l_Lean_Meta_registerGetEqnsFn(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -3450,7 +3640,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2; +x_5 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2; x_6 = l_Lean_Meta_registerGetUnfoldEqnFn(x_5, x_4); if (lean_obj_tag(x_6) == 0) { @@ -3460,7 +3650,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__5; x_9 = 0; -x_10 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19; +x_10 = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19; x_11 = l_Lean_registerTraceClass(x_8, x_9, x_10, x_7); return x_11; } @@ -3587,6 +3777,10 @@ l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkPr lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__17); l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18 = _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18(); lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__18); +l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19 = _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19(); +lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__19); +l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20 = _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20(); +lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___lambda__1___closed__20); l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1 = _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1(); lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__1); l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__2 = _init_l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof_go___closed__2(); @@ -3607,64 +3801,68 @@ l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkPr lean_mark_persistent(l___private_Lean_Elab_PreDefinition_Structural_Eqns_0__Lean_Elab_Structural_mkProof___closed__2); l_Lean_Elab_Structural_mkEqns___lambda__1___closed__1 = _init_l_Lean_Elab_Structural_mkEqns___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Structural_mkEqns___lambda__1___closed__1); +l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1 = _init_l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_Structural_mkEqns___lambda__2___closed__1); l_Lean_Elab_Structural_mkEqns___closed__1 = _init_l_Lean_Elab_Structural_mkEqns___closed__1(); lean_mark_persistent(l_Lean_Elab_Structural_mkEqns___closed__1); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__1); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__2); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__3); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194____closed__4); -if (builtin) {res = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1194_(lean_io_mk_world()); +l_Lean_Elab_Structural_mkEqns___closed__2 = _init_l_Lean_Elab_Structural_mkEqns___closed__2(); +lean_mark_persistent(l_Lean_Elab_Structural_mkEqns___closed__2); +l_Lean_Elab_Structural_mkEqns___closed__3 = _init_l_Lean_Elab_Structural_mkEqns___closed__3(); +lean_mark_persistent(l_Lean_Elab_Structural_mkEqns___closed__3); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__1); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__2); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__3); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__4 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__4(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195____closed__4); +if (builtin) {res = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1195_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Structural_eqnInfoExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Structural_eqnInfoExt); lean_dec_ref(res); }l_Lean_Elab_Structural_registerEqnsInfo___closed__1 = _init_l_Lean_Elab_Structural_registerEqnsInfo___closed__1(); lean_mark_persistent(l_Lean_Elab_Structural_registerEqnsInfo___closed__1); -l_Lean_Elab_Structural_registerEqnsInfo___closed__2 = _init_l_Lean_Elab_Structural_registerEqnsInfo___closed__2(); -lean_mark_persistent(l_Lean_Elab_Structural_registerEqnsInfo___closed__2); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__1); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__2); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__3); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__4); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__5); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__6); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__7); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__8); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__9); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__10); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__11); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__12); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__13); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__14); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__15); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__16); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__17); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__18); -l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19(); -lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437____closed__19); -if (builtin) {res = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1437_(lean_io_mk_world()); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__1); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__2); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__3); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__4); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__5); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__6); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__7); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__8); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__9); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__10); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__11); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__12); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__13); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__14); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__15); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__16); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__17); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__18); +l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19 = _init_l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19(); +lean_mark_persistent(l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438____closed__19); +if (builtin) {res = l_Lean_Elab_Structural_initFn____x40_Lean_Elab_PreDefinition_Structural_Eqns___hyg_1438_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c index ab726eb7c392..31e540f08516 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/IndPred.c @@ -77,6 +77,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___lambda__3___bo LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___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_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___lambda__5___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___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*); +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___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*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); @@ -89,7 +90,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at___private_Lean_Elab_PreDef static lean_object* l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___closed__3; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___lambda__2___closed__1; -uint8_t l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_mkIndPredBRecOn___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_PreDefinition_Structural_IndPred_0__Lean_Elab_Structural_replaceIndPredRecApps_loop___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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*); @@ -4760,7 +4760,7 @@ if (x_9 == 0) { lean_object* x_12; uint8_t x_13; x_12 = lean_ctor_get(x_1, 7); -x_13 = l_Array_contains___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__1(x_12, x_8); +x_13 = l_Array_contains___at_Lean_Meta_setMVarUserNamesAt___spec__1(x_12, x_8); if (x_13 == 0) { lean_object* x_14; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/SmartUnfolding.c b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/SmartUnfolding.c index 88d6925dfa56..3dfebc7ab564 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/Structural/SmartUnfolding.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/Structural/SmartUnfolding.c @@ -13,22 +13,24 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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_Structural_addSmartUnfoldingDefAux_visit___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_withEnableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_markSmartUnfoldingMatchAlt(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8(lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mdata___override(lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,59 +45,57 @@ lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_reduceMatcher_x3f___s extern lean_object* l_Lean_Meta_smartUnfoldingSuffix; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__2___closed__2; size_t lean_usize_of_nat(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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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_List_forIn_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__3___closed__2; 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*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___closed__1; lean_object* l_Lean_Meta_markSmartUnfoldingMatch(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10(lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); uint8_t l_Lean_Elab_Structural_recArgHasLooseBVarsAt(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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_st_ref_get(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___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_Structural_addSmartUnfoldingDefAux_visit___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x3f___spec__1(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1(lean_object*); extern lean_object* l_Lean_levelZero; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5; +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_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___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_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; -LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___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_Structural_addSmartUnfoldingDefAux_visit___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__3___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2; lean_object* l_Lean_indentExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8(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* 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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_addSmartUnfoldingDef___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_Structural_addSmartUnfoldingDefAux_visit___spec__3___closed__1; @@ -104,11 +104,7 @@ lean_object* l_Lean_Meta_MatcherApp_toExpr(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2; -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(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* l_Array_ofSubarray___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1___boxed(lean_object*); size_t lean_usize_add(size_t, size_t); lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); @@ -119,6 +115,7 @@ lean_object* l_List_redLength___rarg(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isAuxRecursorWithSuffix(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); 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*); @@ -126,15 +123,15 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__3___closed__3; static lean_object* l_panic___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__2___closed__1; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1___boxed(lean_object*); lean_object* l_Lean_InductiveVal_numCtors(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Structural_addSmartUnfoldingDefAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1; lean_object* lean_expr_instantiate1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__3(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_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1; 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* _init_l_panic___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__2___closed__1() { _start: @@ -1774,67 +1771,7 @@ return x_38; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = 0; -x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(x_1, x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_inc(x_12); -lean_dec(x_10); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_10); -if (x_15 == 0) -{ -return x_10; -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_10, 0); -x_17 = lean_ctor_get(x_10, 1); -lean_inc(x_17); -lean_inc(x_16); -lean_dec(x_10); -x_18 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -return x_18; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg___boxed), 8, 0); -return x_2; -} -} -LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1(lean_object* x_1) { _start: { lean_object* x_2; @@ -1854,15 +1791,15 @@ return x_4; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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; @@ -1879,7 +1816,7 @@ lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1; +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1; lean_inc(x_13); x_16 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x3f___spec__1(x_15, x_13); if (lean_obj_tag(x_16) == 0) @@ -1988,7 +1925,7 @@ return x_42; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -1996,16 +1933,16 @@ x_1 = lean_mk_string_from_bytes("unexpected matcher application alternative", 42 return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3() { _start: { lean_object* x_1; @@ -2013,16 +1950,16 @@ x_1 = lean_mk_string_from_bytes("\nat application", 15); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5() { _start: { lean_object* x_1; @@ -2030,16 +1967,16 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { lean_object* x_13; uint8_t x_14; @@ -2055,11 +1992,11 @@ lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); x_15 = l_Lean_indentExpr(x_4); -x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2; +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2; x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_16); lean_ctor_set(x_17, 1, x_15); -x_18 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4; +x_18 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4; x_19 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -2067,7 +2004,7 @@ x_20 = l_Lean_indentExpr(x_5); 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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6; +x_22 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6; x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -2101,12 +2038,12 @@ lean_object* x_29; lean_object* x_30; lean_dec(x_5); lean_dec(x_4); x_29 = lean_box(0); -x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2(x_1, x_2, x_7, x_6, x_3, x_29, x_8, x_9, x_10, x_11, x_12); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2(x_1, x_2, x_7, x_6, x_3, x_29, x_8, x_9, x_10, x_11, x_12); return x_30; } } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_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_Structural_addSmartUnfoldingDefAux_visit___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -2184,7 +2121,7 @@ lean_inc(x_3); lean_inc(x_15); lean_inc(x_2); lean_inc(x_1); -x_31 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3), 12, 5); +x_31 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3), 12, 5); lean_closure_set(x_31, 0, x_1); lean_closure_set(x_31, 1, x_2); lean_closure_set(x_31, 2, x_28); @@ -2195,7 +2132,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_33 = l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(x_15, x_31, x_32, x_8, x_9, x_10, x_11, x_12); +x_33 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_15, x_31, x_32, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_33) == 0) { lean_object* x_34; lean_object* x_35; lean_object* x_36; size_t x_37; size_t x_38; @@ -2261,7 +2198,7 @@ lean_inc(x_3); lean_inc(x_15); lean_inc(x_2); lean_inc(x_1); -x_48 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3), 12, 5); +x_48 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3), 12, 5); lean_closure_set(x_48, 0, x_1); lean_closure_set(x_48, 1, x_2); lean_closure_set(x_48, 2, x_44); @@ -2272,7 +2209,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_50 = l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(x_15, x_48, x_49, x_8, x_9, x_10, x_11, x_12); +x_50 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_15, x_48, x_49, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_50) == 0) { lean_object* x_51; lean_object* x_52; lean_object* x_53; size_t x_54; size_t x_55; @@ -2392,7 +2329,7 @@ lean_inc(x_3); lean_inc(x_15); lean_inc(x_2); lean_inc(x_1); -x_74 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3), 12, 5); +x_74 = lean_alloc_closure((void*)(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3), 12, 5); lean_closure_set(x_74, 0, x_1); lean_closure_set(x_74, 1, x_2); lean_closure_set(x_74, 2, x_70); @@ -2403,7 +2340,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_76 = l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(x_15, x_74, x_75, x_8, x_9, x_10, x_11, x_12); +x_76 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_15, x_74, x_75, x_8, x_9, x_10, x_11, x_12); if (lean_obj_tag(x_76) == 0) { lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; size_t x_81; size_t x_82; @@ -2461,7 +2398,7 @@ return x_87; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___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_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___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) { _start: { lean_object* x_11; @@ -2512,11 +2449,11 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___rarg___boxed), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___rarg___boxed), 10, 0); return x_2; } } @@ -2784,7 +2721,7 @@ x_38 = lean_array_get_size(x_35); x_39 = lean_usize_of_nat(x_38); lean_dec(x_38); x_40 = 0; -x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(x_1, x_2, x_3, x_35, x_39, x_40, x_37, x_4, x_5, x_6, x_7, x_20); +x_41 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8(x_1, x_2, x_3, x_35, x_39, x_40, x_37, x_4, x_5, x_6, x_7, x_20); lean_dec(x_35); if (lean_obj_tag(x_41) == 0) { @@ -2941,7 +2878,7 @@ x_79 = lean_alloc_closure((void*)(l_Lean_Elab_Structural_addSmartUnfoldingDefAux lean_closure_set(x_79, 0, x_1); lean_closure_set(x_79, 1, x_2); x_80 = 0; -x_81 = l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(x_3, x_79, x_80, x_4, x_5, x_6, x_7, x_8); +x_81 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_3, x_79, x_80, x_4, x_5, x_6, x_7, x_8); return x_81; } case 7: @@ -2985,7 +2922,7 @@ lean_closure_set(x_91, 0, x_87); lean_closure_set(x_91, 1, x_1); lean_closure_set(x_91, 2, x_2); x_92 = 0; -x_93 = l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___rarg(x_84, x_85, x_89, x_91, x_92, x_4, x_5, x_6, x_7, x_90); +x_93 = l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___rarg(x_84, x_85, x_89, x_91, x_92, x_4, x_5, x_6, x_7, x_90); return x_93; } else @@ -3240,36 +3177,26 @@ x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_ return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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_3); -lean_dec(x_3); -x_10 = l_Lean_Meta_lambdaTelescope___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___rarg(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__1(x_1); +x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__1(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___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_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); return x_12; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -3277,18 +3204,18 @@ x_13 = lean_unbox_usize(x_5); lean_dec(x_5); x_14 = lean_unbox_usize(x_6); lean_dec(x_6); -x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8(x_1, x_2, x_3, x_4, x_13, x_14, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_4); return x_15; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___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_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_5); lean_dec(x_5); -x_12 = l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__10___rarg(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_withLetDecl___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___rarg(x_1, x_2, x_3, x_4, x_11, x_6, x_7, x_8, x_9, x_10); return x_12; } } @@ -3685,20 +3612,20 @@ l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAu lean_mark_persistent(l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___lambda__3___closed__1); l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___closed__1 = _init_l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___closed__1(); lean_mark_persistent(l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__1___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__2___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__1); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__3); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__4); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__5); -l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__9___lambda__3___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__2___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___spec__8___lambda__3___closed__6); l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___lambda__3___closed__1 = _init_l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Elab_Structural_addSmartUnfoldingDefAux_visit___lambda__3___closed__1); l_Lean_Elab_Structural_addSmartUnfoldingDefAux___closed__1 = _init_l_Lean_Elab_Structural_addSmartUnfoldingDefAux___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c index ae29fc30b470..a9b9715de7b3 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -22,6 +22,7 @@ lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_Simp_simpMatch___spec 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*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo; @@ -38,30 +39,32 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__1___boxed(lean lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_Simp_discharge_x3f_x27___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_mkEqns___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1___closed__1; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__1; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__3; 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_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372_(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__2; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__2___closed__1; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__11; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9; lean_object* l_Lean_MVarId_getType_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__6; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__10; @@ -69,7 +72,6 @@ lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_main(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEq; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3; extern lean_object* l_Lean_casesOnSuffix; lean_object* l_Lean_Meta_Match_MatcherInfo_arity(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___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*); @@ -86,19 +88,21 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__8; lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(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*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_eqnInfoExt; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__3___boxed(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*); @@ -107,13 +111,15 @@ static uint32_t l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkPr static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___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_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__4; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10; lean_object* lean_st_ref_take(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix(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_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__4; @@ -123,32 +129,29 @@ static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_DefKind_isTheorem(uint8_t); lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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_WF_simpMatchWF_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3; lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__6; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__6; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__4; lean_object* l_Lean_Meta_mkEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_object*); lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,11 +164,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Ela static lean_object* l_Lean_Elab_WF_mkEqns___closed__1; extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_registerEqnsInfo(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_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__1; @@ -174,17 +177,22 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__3; lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__1(lean_object*, size_t, size_t); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_mkEqns___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f_pre(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_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__5; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4; +extern lean_object* l_Lean_diagnostics; +static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19; LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); lean_object* l_Lean_Elab_Eqns_simpIf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; -lean_object* l_Lean_Elab_Eqns_mkEqnTypes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Elab_Eqns_mkEqnTypes(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__4; lean_object* l_Lean_Meta_splitTarget_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__1; @@ -193,6 +201,8 @@ static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instBEqOrigin___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; lean_object* l_Lean_Meta_SplitIf_mkDischarge_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Split_getSimpMatchContext___rarg(lean_object*, lean_object*); @@ -200,14 +210,15 @@ static lean_object* l_panic___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__2___cl static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__3___closed__1; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__12; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6; extern lean_object* l_Lean_Meta_eqnThmSuffixBase; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___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_name_append_index_after(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__6; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__3; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12; +static lean_object* l_Lean_Elab_WF_mkEqns___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_registerEqnsInfo___spec__2(size_t, size_t, lean_object*); lean_object* l_Lean_Expr_constLevels_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___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*); @@ -217,30 +228,30 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__3___boxed(lean LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__2; static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__14; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__8; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__9; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___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_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__10; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102_(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__3; static lean_object* l_panic___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__2___closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__7; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__5; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__1; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__5; lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -248,6 +259,7 @@ static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__15; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__4; extern lean_object* l_Lean_Expr_instHashable; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__1; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__1; @@ -256,27 +268,25 @@ lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__6; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18; lean_object* l_Lean_Meta_ensureEqnReservedNamesAvailable(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__4; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__2; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__2; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13; -lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__9; -static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__16; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2; lean_object* l_Lean_Elab_Eqns_whnfReducibleLHS_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_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__1; size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_instHashableOrigin___boxed(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f_pre___closed__1; lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Meta_casesOnStuckLHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -301,6 +311,7 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; lean_object* l_Lean_MVarId_intros(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -309,13 +320,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns(lean_object*, lean_object*, lean_ lean_object* l_Lean_Meta_applySimpResultToTarget(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_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__5; lean_object* l_Lean_Elab_Eqns_deltaLHS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__2; lean_object* l_Lean_InductiveVal_numCtors(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14; static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__3(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_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Meta_subst_substEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2721,12 +2735,49 @@ return x_3; static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_2 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_2 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8; +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_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_simpMatchWF_x3f_pre), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8() { +static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11() { _start: { lean_object* x_1; @@ -2734,7 +2785,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__1___bo return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9() { +static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12() { _start: { lean_object* x_1; @@ -2742,7 +2793,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__2___bo return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10() { +static lean_object* _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13() { _start: { lean_object* x_1; @@ -2779,7 +2830,7 @@ lean_inc(x_1); x_14 = l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(x_1, x_13, x_2, x_3, x_4, x_5, x_12); if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); @@ -2791,203 +2842,205 @@ lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; -x_21 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8; -x_22 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9; -x_23 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10; -x_24 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_21); -lean_ctor_set(x_24, 2, x_22); -lean_ctor_set(x_24, 3, x_23); -lean_ctor_set(x_24, 4, x_15); -x_25 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_20 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10; +x_21 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11; +x_22 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12; +x_23 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13; +x_24 = 1; +x_25 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_25, 0, x_20); +lean_ctor_set(x_25, 1, x_21); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_23); +lean_ctor_set(x_25, 4, x_15); +lean_ctor_set_uint8(x_25, sizeof(void*)*5, x_24); +x_26 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9; lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_11); -x_26 = l_Lean_Meta_Simp_main(x_11, x_18, x_25, x_24, x_2, x_3, x_4, x_5, x_19); -if (lean_obj_tag(x_26) == 0) +x_27 = l_Lean_Meta_Simp_main(x_11, x_18, x_26, x_25, x_2, x_3, x_4, x_5, x_19); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_27, 0); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); lean_dec(x_27); +x_30 = lean_ctor_get(x_28, 0); +lean_inc(x_30); +lean_dec(x_28); lean_inc(x_1); -x_30 = l_Lean_Meta_applySimpResultToTarget(x_1, x_11, x_29, x_2, x_3, x_4, x_5, x_28); +x_31 = l_Lean_Meta_applySimpResultToTarget(x_1, x_11, x_30, x_2, x_3, x_4, x_5, x_29); lean_dec(x_11); -if (lean_obj_tag(x_30) == 0) +if (lean_obj_tag(x_31) == 0) { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_30); -if (x_31 == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) { -lean_object* x_32; uint8_t x_33; -x_32 = lean_ctor_get(x_30, 0); -x_33 = lean_name_eq(x_1, x_32); +lean_object* x_33; uint8_t x_34; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_name_eq(x_1, x_33); lean_dec(x_1); -if (x_33 == 0) +if (x_34 == 0) { -lean_object* x_34; -x_34 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_30, 0, x_34); -return x_30; +lean_object* x_35; +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_31, 0, x_35); +return x_31; } else { -lean_object* x_35; -lean_dec(x_32); -x_35 = lean_box(0); -lean_ctor_set(x_30, 0, x_35); -return x_30; +lean_object* x_36; +lean_dec(x_33); +x_36 = lean_box(0); +lean_ctor_set(x_31, 0, x_36); +return x_31; } } else { -lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_36 = lean_ctor_get(x_30, 0); -x_37 = lean_ctor_get(x_30, 1); +lean_object* x_37; lean_object* x_38; uint8_t 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_inc(x_36); -lean_dec(x_30); -x_38 = lean_name_eq(x_1, x_36); +lean_dec(x_31); +x_39 = lean_name_eq(x_1, x_37); lean_dec(x_1); -if (x_38 == 0) +if (x_39 == 0) { -lean_object* x_39; lean_object* x_40; -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, 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_40; lean_object* x_41; +x_40 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_40, 0, x_37); +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_41; lean_object* x_42; -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_37); -return x_42; +lean_object* x_42; lean_object* x_43; +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_38); +return x_43; } } } else { -uint8_t x_43; +uint8_t x_44; lean_dec(x_1); -x_43 = !lean_is_exclusive(x_30); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_31); +if (x_44 == 0) { -return x_30; +return x_31; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_30, 0); -x_45 = lean_ctor_get(x_30, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_31, 0); +x_46 = lean_ctor_get(x_31, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_30); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_dec(x_31); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -uint8_t x_47; +uint8_t x_48; lean_dec(x_11); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_47 = !lean_is_exclusive(x_26); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_27); +if (x_48 == 0) { -return x_26; +return x_27; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_26, 0); -x_49 = lean_ctor_get(x_26, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_27, 0); +x_50 = lean_ctor_get(x_27, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_26); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_dec(x_27); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -uint8_t x_51; +uint8_t x_52; lean_dec(x_11); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_51 = !lean_is_exclusive(x_14); -if (x_51 == 0) +x_52 = !lean_is_exclusive(x_14); +if (x_52 == 0) { return x_14; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_14, 0); -x_53 = lean_ctor_get(x_14, 1); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_14, 0); +x_54 = lean_ctor_get(x_14, 1); +lean_inc(x_54); lean_inc(x_53); -lean_inc(x_52); 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); -return x_54; +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } else { -uint8_t x_55; +uint8_t x_56; lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = !lean_is_exclusive(x_7); -if (x_55 == 0) +x_56 = !lean_is_exclusive(x_7); +if (x_56 == 0) { return x_7; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_7, 0); -x_57 = lean_ctor_get(x_7, 1); +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_7, 0); +x_58 = lean_ctor_get(x_7, 1); +lean_inc(x_58); lean_inc(x_57); -lean_inc(x_56); lean_dec(x_7); -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; +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; } } } @@ -3681,22 +3734,10 @@ return x_2; static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__3; -x_2 = lean_unsigned_to_nat(0u); -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5() { -_start: -{ uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = 1; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__3; -x_3 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4; +x_3 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; x_4 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_4, 0, x_2); lean_ctor_set(x_4, 1, x_3); @@ -3704,39 +3745,66 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_1); return x_4; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5() { _start: { -lean_object* x_1; uint32_t x_2; lean_object* x_3; uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +lean_object* x_1; uint32_t x_2; lean_object* x_3; uint32_t x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; x_1 = lean_box(0); x_2 = 0; x_3 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__1; x_4 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__2; x_5 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; -x_6 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; -x_7 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_7, 0, x_3); -lean_ctor_set(x_7, 1, x_5); -lean_ctor_set(x_7, 2, x_6); -lean_ctor_set(x_7, 3, x_1); -lean_ctor_set_uint32(x_7, sizeof(void*)*4, x_4); -lean_ctor_set_uint32(x_7, sizeof(void*)*4 + 4, x_2); -return x_7; +x_6 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4; +x_7 = lean_unsigned_to_nat(0u); +x_8 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_8, 0, x_3); +lean_ctor_set(x_8, 1, x_5); +lean_ctor_set(x_8, 2, x_6); +lean_ctor_set(x_8, 3, x_1); +lean_ctor_set(x_8, 4, x_7); +lean_ctor_set_uint32(x_8, sizeof(void*)*5, x_4); +lean_ctor_set_uint32(x_8, sizeof(void*)*5 + 4, x_2); +return x_8; +} +} +static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_2 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; } } static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("'\n", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8() { +static lean_object* _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; +x_1 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -3836,9 +3904,9 @@ x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); lean_dec(x_23); x_26 = lean_box(0); -x_27 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6; +x_27 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; x_28 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; -x_29 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__6; +x_29 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); @@ -3932,7 +4000,7 @@ x_48 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefi x_49 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_49, 0, x_48); lean_ctor_set(x_49, 1, x_47); -x_50 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8; +x_50 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___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); @@ -5266,7 +5334,7 @@ lean_inc(x_6); x_16 = l_Lean_Meta_mkEq(x_15, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); @@ -5284,191 +5352,427 @@ x_23 = lean_ctor_get(x_3, 1); lean_inc(x_23); lean_dec(x_3); x_24 = l_Lean_Expr_mvarId_x21(x_21); -x_25 = l_Lean_Elab_Eqns_mkEqnTypes(x_23, x_24, x_6, x_7, x_8, x_9, x_22); -return x_25; +x_25 = !lean_is_exclusive(x_6); +if (x_25 == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_6, 0); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_28 = 2; +lean_ctor_set_uint8(x_26, 9, x_28); +x_29 = 0; +x_30 = l_Lean_Elab_Eqns_mkEqnTypes(x_29, x_23, x_24, x_6, x_7, x_8, x_9, x_22); +if (lean_obj_tag(x_30) == 0) +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_30); +if (x_31 == 0) +{ +return x_30; } else { -uint8_t x_26; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_3); -x_26 = !lean_is_exclusive(x_16); -if (x_26 == 0) +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_30, 0); +x_33 = lean_ctor_get(x_30, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_30); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; +} +} +else { -return x_16; +uint8_t x_35; +x_35 = !lean_is_exclusive(x_30); +if (x_35 == 0) +{ +return x_30; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_16, 0); -x_28 = lean_ctor_get(x_16, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_16); -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_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_30, 0); +x_37 = lean_ctor_get(x_30, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_30); +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 +{ +uint8_t x_39; uint8_t x_40; uint8_t x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; +x_39 = lean_ctor_get_uint8(x_26, 0); +x_40 = lean_ctor_get_uint8(x_26, 1); +x_41 = lean_ctor_get_uint8(x_26, 2); +x_42 = lean_ctor_get_uint8(x_26, 3); +x_43 = lean_ctor_get_uint8(x_26, 4); +x_44 = lean_ctor_get_uint8(x_26, 5); +x_45 = lean_ctor_get_uint8(x_26, 6); +x_46 = lean_ctor_get_uint8(x_26, 7); +x_47 = lean_ctor_get_uint8(x_26, 8); +x_48 = lean_ctor_get_uint8(x_26, 10); +x_49 = lean_ctor_get_uint8(x_26, 11); +x_50 = lean_ctor_get_uint8(x_26, 12); +lean_dec(x_26); +x_51 = 2; +x_52 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_52, 0, x_39); +lean_ctor_set_uint8(x_52, 1, x_40); +lean_ctor_set_uint8(x_52, 2, x_41); +lean_ctor_set_uint8(x_52, 3, x_42); +lean_ctor_set_uint8(x_52, 4, x_43); +lean_ctor_set_uint8(x_52, 5, x_44); +lean_ctor_set_uint8(x_52, 6, x_45); +lean_ctor_set_uint8(x_52, 7, x_46); +lean_ctor_set_uint8(x_52, 8, x_47); +lean_ctor_set_uint8(x_52, 9, x_51); +lean_ctor_set_uint8(x_52, 10, x_48); +lean_ctor_set_uint8(x_52, 11, x_49); +lean_ctor_set_uint8(x_52, 12, x_50); +lean_ctor_set(x_6, 0, x_52); +x_53 = 0; +x_54 = l_Lean_Elab_Eqns_mkEqnTypes(x_53, x_23, x_24, x_6, x_7, x_8, x_9, x_22); +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, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_57 = x_54; +} else { + lean_dec_ref(x_54); + x_57 = lean_box(0); +} +if (lean_is_scalar(x_57)) { + x_58 = lean_alloc_ctor(0, 2, 0); +} else { + x_58 = x_57; +} +lean_ctor_set(x_58, 0, x_55); +lean_ctor_set(x_58, 1, x_56); +return x_58; } -static lean_object* _init_l_Lean_Elab_WF_mkEqns___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_Meta_tactic_hygienic; -return x_1; +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_59 = lean_ctor_get(x_54, 0); +lean_inc(x_59); +x_60 = lean_ctor_get(x_54, 1); +lean_inc(x_60); +if (lean_is_exclusive(x_54)) { + lean_ctor_release(x_54, 0); + lean_ctor_release(x_54, 1); + x_61 = x_54; +} else { + lean_dec_ref(x_54); + x_61 = lean_box(0); } +if (lean_is_scalar(x_61)) { + x_62 = lean_alloc_ctor(1, 2, 0); +} else { + x_62 = x_61; } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns(lean_object* x_1, 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_ctor_set(x_62, 0, x_59); +lean_ctor_set(x_62, 1, x_60); +return x_62; +} +} +} +else { -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; -x_8 = lean_ctor_get(x_2, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_8, 3); -lean_inc(x_9); -lean_inc(x_1); -lean_inc(x_8); -x_10 = lean_alloc_closure((void*)(l_Lean_Elab_WF_mkEqns___lambda__1), 10, 3); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_1); -lean_closure_set(x_10, 2, x_2); -x_11 = 1; -x_12 = lean_box(x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed), 8, 3); -lean_closure_set(x_13, 0, x_9); -lean_closure_set(x_13, 1, x_10); -lean_closure_set(x_13, 2, x_12); -x_14 = !lean_is_exclusive(x_5); -if (x_14 == 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; uint8_t x_69; uint8_t x_70; uint8_t x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; lean_object* x_88; +x_63 = lean_ctor_get(x_6, 0); +x_64 = lean_ctor_get(x_6, 1); +x_65 = lean_ctor_get(x_6, 2); +x_66 = lean_ctor_get(x_6, 3); +x_67 = lean_ctor_get(x_6, 4); +x_68 = lean_ctor_get(x_6, 5); +x_69 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); +x_70 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); +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_dec(x_6); +x_71 = lean_ctor_get_uint8(x_63, 0); +x_72 = lean_ctor_get_uint8(x_63, 1); +x_73 = lean_ctor_get_uint8(x_63, 2); +x_74 = lean_ctor_get_uint8(x_63, 3); +x_75 = lean_ctor_get_uint8(x_63, 4); +x_76 = lean_ctor_get_uint8(x_63, 5); +x_77 = lean_ctor_get_uint8(x_63, 6); +x_78 = lean_ctor_get_uint8(x_63, 7); +x_79 = lean_ctor_get_uint8(x_63, 8); +x_80 = lean_ctor_get_uint8(x_63, 10); +x_81 = lean_ctor_get_uint8(x_63, 11); +x_82 = lean_ctor_get_uint8(x_63, 12); +if (lean_is_exclusive(x_63)) { + x_83 = x_63; +} else { + lean_dec_ref(x_63); + x_83 = lean_box(0); +} +x_84 = 2; +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 0, 13); +} else { + x_85 = x_83; +} +lean_ctor_set_uint8(x_85, 0, x_71); +lean_ctor_set_uint8(x_85, 1, x_72); +lean_ctor_set_uint8(x_85, 2, x_73); +lean_ctor_set_uint8(x_85, 3, x_74); +lean_ctor_set_uint8(x_85, 4, x_75); +lean_ctor_set_uint8(x_85, 5, x_76); +lean_ctor_set_uint8(x_85, 6, x_77); +lean_ctor_set_uint8(x_85, 7, x_78); +lean_ctor_set_uint8(x_85, 8, x_79); +lean_ctor_set_uint8(x_85, 9, x_84); +lean_ctor_set_uint8(x_85, 10, x_80); +lean_ctor_set_uint8(x_85, 11, x_81); +lean_ctor_set_uint8(x_85, 12, x_82); +x_86 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_86, 0, x_85); +lean_ctor_set(x_86, 1, x_64); +lean_ctor_set(x_86, 2, x_65); +lean_ctor_set(x_86, 3, x_66); +lean_ctor_set(x_86, 4, x_67); +lean_ctor_set(x_86, 5, x_68); +lean_ctor_set_uint8(x_86, sizeof(void*)*6, x_69); +lean_ctor_set_uint8(x_86, sizeof(void*)*6 + 1, x_70); +x_87 = 0; +x_88 = l_Lean_Elab_Eqns_mkEqnTypes(x_87, x_23, x_24, x_86, x_7, x_8, x_9, x_22); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = lean_ctor_get(x_88, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_91 = x_88; +} else { + lean_dec_ref(x_88); + x_91 = lean_box(0); +} +if (lean_is_scalar(x_91)) { + x_92 = lean_alloc_ctor(0, 2, 0); +} else { + x_92 = x_91; +} +lean_ctor_set(x_92, 0, x_89); +lean_ctor_set(x_92, 1, x_90); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_88, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_88, 1); +lean_inc(x_94); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_95 = x_88; +} else { + lean_dec_ref(x_88); + x_95 = lean_box(0); +} +if (lean_is_scalar(x_95)) { + x_96 = lean_alloc_ctor(1, 2, 0); +} else { + x_96 = x_95; +} +lean_ctor_set(x_96, 0, x_93); +lean_ctor_set(x_96, 1, x_94); +return x_96; +} +} +} +else +{ +uint8_t x_97; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_3); +x_97 = !lean_is_exclusive(x_16); +if (x_97 == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_5, 2); -x_16 = l_Lean_Elab_WF_mkEqns___closed__1; +return x_16; +} +else +{ +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_16, 0); +x_99 = lean_ctor_get(x_16, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_16); +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; +} +} +} +} +static lean_object* _init_l_Lean_Elab_WF_mkEqns___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +x_13 = lean_ctor_get(x_9, 4); +lean_dec(x_13); +x_14 = lean_ctor_get(x_9, 2); +lean_dec(x_14); +x_15 = l_Lean_Elab_WF_mkEqns___lambda__2___closed__1; +x_16 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_15); +lean_ctor_set(x_9, 4, x_16); +lean_ctor_set(x_9, 2, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*11 + 1, x_2); x_17 = 0; -x_18 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_15, x_16, x_17); -lean_ctor_set(x_5, 2, x_18); -lean_inc(x_6); +lean_inc(x_10); +lean_inc(x_9); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_19 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_13, x_17, x_3, x_4, x_5, x_6, x_7); -if (lean_obj_tag(x_19) == 0) +x_18 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_3, x_17, x_4, x_5, x_9, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_20 = lean_ctor_get(x_19, 0); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); +lean_dec(x_18); +x_21 = lean_box(0); +x_22 = lean_array_get_size(x_19); +x_23 = lean_unsigned_to_nat(0u); +x_24 = lean_unsigned_to_nat(1u); +x_25 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; +lean_inc(x_22); +x_26 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1(x_6, x_7, x_19, x_21, x_22, x_22, x_23, x_22, x_24, x_25, x_4, x_5, x_9, x_10, x_20); +lean_dec(x_22); lean_dec(x_19); -x_22 = lean_box(0); -x_23 = lean_array_get_size(x_20); -x_24 = lean_unsigned_to_nat(0u); -x_25 = lean_unsigned_to_nat(1u); -x_26 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; -lean_inc(x_23); -x_27 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1(x_1, x_8, x_20, x_22, x_23, x_23, x_24, x_23, x_25, x_26, x_3, x_4, x_5, x_6, x_21); -lean_dec(x_23); -lean_dec(x_20); -lean_dec(x_8); -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_26) == 0) { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -return x_27; +return x_26; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = lean_ctor_get(x_27, 0); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); lean_inc(x_29); -lean_dec(x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +uint8_t x_31; +x_31 = !lean_is_exclusive(x_26); +if (x_31 == 0) { -return x_27; +return x_26; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_26, 0); +x_33 = lean_ctor_get(x_26, 1); lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_inc(x_32); +lean_dec(x_26); +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_36; -lean_dec(x_5); -lean_dec(x_8); +uint8_t x_35; +lean_dec(x_9); +lean_dec(x_10); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_36 = !lean_is_exclusive(x_19); -if (x_36 == 0) +x_35 = !lean_is_exclusive(x_18); +if (x_35 == 0) { -return x_19; +return x_18; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_19, 0); -x_38 = lean_ctor_get(x_19, 1); -lean_inc(x_38); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_18, 0); +x_37 = lean_ctor_get(x_18, 1); lean_inc(x_37); -lean_dec(x_19); -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_inc(x_36); +lean_dec(x_18); +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_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; uint8_t x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_40 = lean_ctor_get(x_5, 0); -x_41 = lean_ctor_get(x_5, 1); -x_42 = lean_ctor_get(x_5, 2); -x_43 = lean_ctor_get(x_5, 3); -x_44 = lean_ctor_get(x_5, 4); -x_45 = lean_ctor_get(x_5, 5); -x_46 = lean_ctor_get(x_5, 6); -x_47 = lean_ctor_get(x_5, 7); -x_48 = lean_ctor_get(x_5, 8); -x_49 = lean_ctor_get(x_5, 9); -x_50 = lean_ctor_get(x_5, 10); -x_51 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); -x_52 = lean_ctor_get_uint8(x_5, sizeof(void*)*11 + 1); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +x_39 = lean_ctor_get(x_9, 0); +x_40 = lean_ctor_get(x_9, 1); +x_41 = lean_ctor_get(x_9, 3); +x_42 = lean_ctor_get(x_9, 5); +x_43 = lean_ctor_get(x_9, 6); +x_44 = lean_ctor_get(x_9, 7); +x_45 = lean_ctor_get(x_9, 8); +x_46 = lean_ctor_get(x_9, 9); +x_47 = lean_ctor_get(x_9, 10); +x_48 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); lean_inc(x_47); lean_inc(x_46); lean_inc(x_45); @@ -5477,64 +5781,88 @@ lean_inc(x_43); lean_inc(x_42); lean_inc(x_41); lean_inc(x_40); -lean_dec(x_5); -x_53 = l_Lean_Elab_WF_mkEqns___closed__1; -x_54 = 0; -x_55 = l_Lean_Option_set___at_Lean_Elab_Eqns_tryURefl___spec__1(x_42, x_53, x_54); -x_56 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_56, 0, x_40); -lean_ctor_set(x_56, 1, x_41); -lean_ctor_set(x_56, 2, x_55); -lean_ctor_set(x_56, 3, x_43); -lean_ctor_set(x_56, 4, x_44); -lean_ctor_set(x_56, 5, x_45); -lean_ctor_set(x_56, 6, x_46); -lean_ctor_set(x_56, 7, x_47); -lean_ctor_set(x_56, 8, x_48); -lean_ctor_set(x_56, 9, x_49); -lean_ctor_set(x_56, 10, x_50); -lean_ctor_set_uint8(x_56, sizeof(void*)*11, x_51); -lean_ctor_set_uint8(x_56, sizeof(void*)*11 + 1, x_52); -lean_inc(x_6); -lean_inc(x_56); +lean_inc(x_39); +lean_dec(x_9); +x_49 = l_Lean_Elab_WF_mkEqns___lambda__2___closed__1; +x_50 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_49); +x_51 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_51, 0, x_39); +lean_ctor_set(x_51, 1, x_40); +lean_ctor_set(x_51, 2, x_1); +lean_ctor_set(x_51, 3, x_41); +lean_ctor_set(x_51, 4, x_50); +lean_ctor_set(x_51, 5, x_42); +lean_ctor_set(x_51, 6, x_43); +lean_ctor_set(x_51, 7, x_44); +lean_ctor_set(x_51, 8, x_45); +lean_ctor_set(x_51, 9, x_46); +lean_ctor_set(x_51, 10, x_47); +lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_48); +lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_2); +x_52 = 0; +lean_inc(x_10); +lean_inc(x_51); +lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_57 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_13, x_54, x_3, x_4, x_56, x_6, x_7); -if (lean_obj_tag(x_57) == 0) +x_53 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_matchesInstance___spec__1___rarg(x_3, x_52, x_4, x_5, x_51, x_10, x_11); +if (lean_obj_tag(x_53) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_box(0); +x_57 = lean_array_get_size(x_54); +x_58 = lean_unsigned_to_nat(0u); +x_59 = lean_unsigned_to_nat(1u); +x_60 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; +lean_inc(x_57); +x_61 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1(x_6, x_7, x_54, x_56, x_57, x_57, x_58, x_57, x_59, x_60, x_4, x_5, x_51, x_10, x_55); lean_dec(x_57); -x_60 = lean_box(0); -x_61 = lean_array_get_size(x_58); -x_62 = lean_unsigned_to_nat(0u); -x_63 = lean_unsigned_to_nat(1u); -x_64 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; -lean_inc(x_61); -x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1(x_1, x_8, x_58, x_60, x_61, x_61, x_62, x_61, x_63, x_64, x_3, x_4, x_56, x_6, x_59); -lean_dec(x_61); -lean_dec(x_58); -lean_dec(x_8); -if (lean_obj_tag(x_65) == 0) +lean_dec(x_54); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_64 = x_61; +} else { + lean_dec_ref(x_61); + x_64 = lean_box(0); +} +if (lean_is_scalar(x_64)) { + x_65 = lean_alloc_ctor(0, 2, 0); +} else { + x_65 = x_64; +} +lean_ctor_set(x_65, 0, x_62); +lean_ctor_set(x_65, 1, x_63); +return x_65; +} +else { lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_65, 0); +x_66 = lean_ctor_get(x_61, 0); lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); +x_67 = lean_ctor_get(x_61, 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; +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_68 = x_61; } else { - lean_dec_ref(x_65); + lean_dec_ref(x_61); x_68 = lean_box(0); } if (lean_is_scalar(x_68)) { - x_69 = lean_alloc_ctor(0, 2, 0); + x_69 = lean_alloc_ctor(1, 2, 0); } else { x_69 = x_68; } @@ -5542,19 +5870,25 @@ 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; -x_70 = lean_ctor_get(x_65, 0); +lean_dec(x_51); +lean_dec(x_10); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_70 = lean_ctor_get(x_53, 0); lean_inc(x_70); -x_71 = lean_ctor_get(x_65, 1); +x_71 = lean_ctor_get(x_53, 1); lean_inc(x_71); -if (lean_is_exclusive(x_65)) { - lean_ctor_release(x_65, 0); - lean_ctor_release(x_65, 1); - x_72 = x_65; +if (lean_is_exclusive(x_53)) { + lean_ctor_release(x_53, 0); + lean_ctor_release(x_53, 1); + x_72 = x_53; } else { - lean_dec_ref(x_65); + lean_dec_ref(x_53); x_72 = lean_box(0); } if (lean_is_scalar(x_72)) { @@ -5567,35 +5901,173 @@ lean_ctor_set(x_73, 1, x_71); return x_73; } } +} +} +static lean_object* _init_l_Lean_Elab_WF_mkEqns___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_tactic_hygienic; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_mkEqns___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_WF_mkEqns___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__7; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t 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; uint8_t x_24; uint8_t x_25; +x_8 = lean_ctor_get(x_2, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 3); +lean_inc(x_9); +lean_inc(x_1); +lean_inc(x_8); +x_10 = lean_alloc_closure((void*)(l_Lean_Elab_WF_mkEqns___lambda__1), 10, 3); +lean_closure_set(x_10, 0, x_8); +lean_closure_set(x_10, 1, x_1); +lean_closure_set(x_10, 2, x_2); +x_11 = 1; +x_12 = lean_box(x_11); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed), 8, 3); +lean_closure_set(x_13, 0, x_9); +lean_closure_set(x_13, 1, x_10); +lean_closure_set(x_13, 2, x_12); +x_14 = lean_ctor_get(x_5, 2); +lean_inc(x_14); +x_15 = l_Lean_Elab_WF_mkEqns___closed__1; +x_16 = 0; +x_17 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_14, x_15, x_16); +x_18 = l_Lean_Elab_WF_mkEqns___closed__2; +x_19 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_17, x_18); +x_20 = lean_st_ref_get(x_6, x_7); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_Kernel_isDiagnosticsEnabled(x_23); +lean_dec(x_23); +if (x_24 == 0) +{ +if (x_19 == 0) +{ +x_25 = x_11; +goto block_53; +} +else +{ +x_25 = x_16; +goto block_53; +} +} +else +{ +if (x_19 == 0) +{ +x_25 = x_16; +goto block_53; +} else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_56); +x_25 = x_11; +goto block_53; +} +} +block_53: +{ +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_st_ref_take(x_6, x_22); +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = !lean_is_exclusive(x_27); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_30 = lean_ctor_get(x_27, 0); +x_31 = lean_ctor_get(x_27, 4); +lean_dec(x_31); +x_32 = l_Lean_Kernel_enableDiag(x_30, x_19); +x_33 = l_Lean_Elab_WF_mkEqns___closed__3; +lean_ctor_set(x_27, 4, x_33); +lean_ctor_set(x_27, 0, x_32); +x_34 = lean_st_ref_set(x_6, x_27, x_28); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +lean_dec(x_34); +x_36 = lean_box(0); +x_37 = l_Lean_Elab_WF_mkEqns___lambda__2(x_17, x_19, x_13, x_3, x_4, x_1, x_8, x_36, x_5, x_6, x_35); lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_74 = lean_ctor_get(x_57, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_57, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_76 = x_57; -} else { - lean_dec_ref(x_57); - x_76 = lean_box(0); +return x_37; } -if (lean_is_scalar(x_76)) { - x_77 = lean_alloc_ctor(1, 2, 0); -} else { - x_77 = x_76; +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_38 = lean_ctor_get(x_27, 0); +x_39 = lean_ctor_get(x_27, 1); +x_40 = lean_ctor_get(x_27, 2); +x_41 = lean_ctor_get(x_27, 3); +x_42 = lean_ctor_get(x_27, 5); +x_43 = lean_ctor_get(x_27, 6); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_27); +x_44 = l_Lean_Kernel_enableDiag(x_38, x_19); +x_45 = l_Lean_Elab_WF_mkEqns___closed__3; +x_46 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_39); +lean_ctor_set(x_46, 2, x_40); +lean_ctor_set(x_46, 3, x_41); +lean_ctor_set(x_46, 4, x_45); +lean_ctor_set(x_46, 5, x_42); +lean_ctor_set(x_46, 6, x_43); +x_47 = lean_st_ref_set(x_6, x_46, x_28); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +lean_dec(x_47); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_WF_mkEqns___lambda__2(x_17, x_19, x_13, x_3, x_4, x_1, x_8, x_49, x_5, x_6, x_48); +lean_dec(x_8); +return x_50; } -lean_ctor_set(x_77, 0, x_74); -lean_ctor_set(x_77, 1, x_75); -return x_77; +} +else +{ +lean_object* x_51; lean_object* x_52; +x_51 = lean_box(0); +x_52 = l_Lean_Elab_WF_mkEqns___lambda__2(x_17, x_19, x_13, x_3, x_4, x_1, x_8, x_51, x_5, x_6, x_22); +lean_dec(x_8); +return x_52; } } } @@ -5624,7 +6096,19 @@ lean_dec(x_2); return x_16; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1() { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___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_2); +lean_dec(x_2); +x_13 = l_Lean_Elab_WF_mkEqns___lambda__2(x_1, x_12, 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); +return x_13; +} +} +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1() { _start: { lean_object* x_1; @@ -5632,7 +6116,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2() { _start: { lean_object* x_1; @@ -5640,7 +6124,7 @@ x_1 = lean_mk_string_from_bytes("WF", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3() { _start: { lean_object* x_1; @@ -5648,23 +6132,23 @@ x_1 = lean_mk_string_from_bytes("eqnInfoExt", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____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_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; -x_3 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2; -x_4 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3; +x_3 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; +x_4 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4; x_3 = l_Lean_mkMapDeclarationExtension___rarg(x_2, x_1); return x_3; } @@ -5974,17 +6458,6 @@ return x_23; static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4; -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__2() { -_start: -{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__3; x_2 = lean_unsigned_to_nat(0u); @@ -5994,7 +6467,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__3() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__2() { _start: { lean_object* x_1; @@ -6002,7 +6475,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Meta_InfoCacheKey_instHashable___boxed), return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__4() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6014,7 +6487,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__5() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__4() { _start: { lean_object* x_1; @@ -6022,21 +6495,21 @@ x_1 = lean_alloc_closure((void*)(l_Lean_instBEqLocalInstance___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__6() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__5; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__4; x_2 = lean_alloc_closure((void*)(l_Array_instBEq___rarg___boxed), 3, 1); lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__7() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__6; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__5; x_2 = l_Lean_Expr_instBEq; x_3 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); lean_closure_set(x_3, 0, x_1); @@ -6044,7 +6517,7 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__8() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__7() { _start: { lean_object* x_1; @@ -6052,21 +6525,21 @@ x_1 = lean_alloc_closure((void*)(l_Lean_instHashableLocalInstance___boxed), 1, 0 return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__9() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__8; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__7; 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_Elab_WF_registerEqnsInfo___closed__10() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__9; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__8; x_2 = l_Lean_Expr_instHashable; x_3 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); lean_closure_set(x_3, 0, x_1); @@ -6074,7 +6547,7 @@ lean_closure_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__11() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6086,7 +6559,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__12() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__11() { _start: { lean_object* x_1; lean_object* x_2; @@ -6097,7 +6570,7 @@ lean_closure_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__13() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -6108,7 +6581,7 @@ lean_closure_set(x_2, 1, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__14() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -6120,11 +6593,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__15() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__14; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__13; x_2 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -6133,14 +6606,14 @@ lean_ctor_set(x_2, 3, x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___closed__16() { +static lean_object* _init_l_Lean_Elab_WF_registerEqnsInfo___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_Elab_WF_registerEqnsInfo___closed__2; -x_2 = l_Lean_Elab_WF_registerEqnsInfo___closed__4; -x_3 = l_Lean_Elab_WF_registerEqnsInfo___closed__11; -x_4 = l_Lean_Elab_WF_registerEqnsInfo___closed__15; +x_1 = l_Lean_Elab_WF_registerEqnsInfo___closed__1; +x_2 = l_Lean_Elab_WF_registerEqnsInfo___closed__3; +x_3 = l_Lean_Elab_WF_registerEqnsInfo___closed__10; +x_4 = l_Lean_Elab_WF_registerEqnsInfo___closed__14; x_5 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -6397,7 +6870,7 @@ goto block_63; block_63: { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_35 = l_Lean_Elab_WF_registerEqnsInfo___closed__1; +x_35 = l_Lean_Elab_WF_mkEqns___closed__3; if (lean_is_scalar(x_33)) { x_36 = lean_alloc_ctor(0, 7, 0); } else { @@ -6427,7 +6900,7 @@ if (x_42 == 0) lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; x_43 = lean_ctor_get(x_40, 1); lean_dec(x_43); -x_44 = l_Lean_Elab_WF_registerEqnsInfo___closed__16; +x_44 = l_Lean_Elab_WF_registerEqnsInfo___closed__15; lean_ctor_set(x_40, 1, x_44); x_45 = lean_st_ref_set(x_6, x_40, x_41); lean_dec(x_6); @@ -6466,7 +6939,7 @@ lean_inc(x_54); lean_inc(x_53); lean_inc(x_52); lean_dec(x_40); -x_56 = l_Lean_Elab_WF_registerEqnsInfo___closed__16; +x_56 = l_Lean_Elab_WF_registerEqnsInfo___closed__15; x_57 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_57, 0, x_52); lean_ctor_set(x_57, 1, x_56); @@ -6911,7 +7384,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1() { _start: { lean_object* x_1; @@ -6919,7 +7392,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_getEqnsFor_x3f), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2() { _start: { lean_object* x_1; @@ -6927,37 +7400,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_getUnfoldFor_x3f), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6() { _start: { lean_object* x_1; @@ -6965,17 +7438,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8() { _start: { lean_object* x_1; @@ -6983,37 +7456,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12() { _start: { lean_object* x_1; @@ -7021,27 +7494,27 @@ x_1 = lean_mk_string_from_bytes("PreDefinition", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15() { _start: { lean_object* x_1; @@ -7049,17 +7522,17 @@ x_1 = lean_mk_string_from_bytes("Eqns", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17() { _start: { lean_object* x_1; @@ -7067,31 +7540,31 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18; -x_2 = lean_unsigned_to_nat(2372u); +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18; +x_2 = lean_unsigned_to_nat(2381u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1; x_3 = l_Lean_Meta_registerGetEqnsFn(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -7099,7 +7572,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2; +x_5 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2; x_6 = l_Lean_Meta_registerGetUnfoldEqnFn(x_5, x_4); if (lean_obj_tag(x_6) == 0) { @@ -7109,7 +7582,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__5; x_9 = 0; -x_10 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19; +x_10 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19; x_11 = l_Lean_registerTraceClass(x_8, x_9, x_10, x_7); return x_11; } @@ -7275,6 +7748,12 @@ l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9 = _init_l_Lean_Elab_WF_si lean_mark_persistent(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9); l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10 = _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10(); lean_mark_persistent(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10); +l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11 = _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11(); +lean_mark_persistent(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__11); +l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12 = _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12(); +lean_mark_persistent(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12); +l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13 = _init_l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13(); +lean_mark_persistent(l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13); l_panic___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1___closed__1 = _init_l_panic___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1___closed__1(); lean_mark_persistent(l_panic___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1___closed__1); l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1(); @@ -7322,6 +7801,8 @@ l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda_ lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7); l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8(); lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__8); +l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9(); +lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9); l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__1 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__1(); lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__1); l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__2 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__2(); @@ -7330,17 +7811,23 @@ l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__1 lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__1); l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__2 = _init_l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__2(); lean_mark_persistent(l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__2); +l_Lean_Elab_WF_mkEqns___lambda__2___closed__1 = _init_l_Lean_Elab_WF_mkEqns___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_mkEqns___lambda__2___closed__1); l_Lean_Elab_WF_mkEqns___closed__1 = _init_l_Lean_Elab_WF_mkEqns___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_mkEqns___closed__1); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__1); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__2); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__3); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102____closed__4); -if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2102_(lean_io_mk_world()); +l_Lean_Elab_WF_mkEqns___closed__2 = _init_l_Lean_Elab_WF_mkEqns___closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_mkEqns___closed__2); +l_Lean_Elab_WF_mkEqns___closed__3 = _init_l_Lean_Elab_WF_mkEqns___closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_mkEqns___closed__3); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4); +if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_WF_eqnInfoExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_WF_eqnInfoExt); @@ -7377,47 +7864,45 @@ l_Lean_Elab_WF_registerEqnsInfo___closed__14 = _init_l_Lean_Elab_WF_registerEqns lean_mark_persistent(l_Lean_Elab_WF_registerEqnsInfo___closed__14); l_Lean_Elab_WF_registerEqnsInfo___closed__15 = _init_l_Lean_Elab_WF_registerEqnsInfo___closed__15(); lean_mark_persistent(l_Lean_Elab_WF_registerEqnsInfo___closed__15); -l_Lean_Elab_WF_registerEqnsInfo___closed__16 = _init_l_Lean_Elab_WF_registerEqnsInfo___closed__16(); -lean_mark_persistent(l_Lean_Elab_WF_registerEqnsInfo___closed__16); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__1); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__2); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__3); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__4); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__5); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__6); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__7); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__8); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__9); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__10); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__11); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__12); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__13); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__14); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__15); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__16); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__17); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__18); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372____closed__19); -if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2372_(lean_io_mk_world()); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19); +if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c index 2ce21f9a997d..8fd49b5e430f 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/GuessLex.c @@ -245,6 +245,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_WF_GuessLex 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_Array_mapMUnsafe_map___at_Lean_Elab_WF_GuessLex_explainMutualFailure___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_List_forIn_loop___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__16(lean_object*); +lean_object* l_Lean_logAt___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__7(lean_object*, lean_object*, uint8_t, 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__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*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_WF_GuessLex_withRecApps_loop___spec__11___boxed(lean_object*, lean_object*); @@ -256,6 +257,7 @@ static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__23; LEAN_EXPORT lean_object* l_Lean_Elab_WF_GuessLex_reportTermArgs(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_Subarray_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_simpleMeasures___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*); @@ -399,7 +401,6 @@ 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_object* l_Lean_Expr_appArg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_WF_GuessLex_explainMutualFailure___spec__12___closed__3; -lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); 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___private_Lean_Elab_PreDefinition_WF_GuessLex_0__Lean_Elab_WF_GuessLex_reprGuessLexRel____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_4254____closed__9; static lean_object* l_Lean_Elab_WF_GuessLex_GuessLexRel_toNatRel___closed__1; @@ -782,7 +783,6 @@ 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*); uint8_t l_Lean_Expr_hasFVar(lean_object*); -lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, 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; @@ -1020,7 +1020,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__2; x_3 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__5; x_4 = l_Lean_Elab_WF_GuessLex_initFn____x40_Lean_Elab_PreDefinition_WF_GuessLex___hyg_9____closed__10; -x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_6____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -27637,7 +27637,7 @@ 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_reportDiag___spec__2(x_49, x_54, x_55, x_5, x_6, x_7, x_8, x_48); +x_56 = l_Lean_logAt___at___private_Lean_Meta_Instances_0__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); @@ -27899,7 +27899,7 @@ x_111 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_111, 0, x_109); lean_ctor_set(x_111, 1, x_110); x_112 = 0; -x_113 = l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2(x_106, x_111, x_112, x_5, x_6, x_7, x_8, x_105); +x_113 = l_Lean_logAt___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__7(x_106, x_111, x_112, x_5, x_6, x_7, x_8, x_105); lean_dec(x_106); x_114 = lean_ctor_get(x_113, 0); lean_inc(x_114); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c index a8319a842eff..e8f67b41a7eb 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/PackMutual.c @@ -18,6 +18,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_packMutual___lambda__2___boxed(lean_obje extern lean_object* l_Lean_Elab_instInhabitedPreDefinition; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_setMVarUserNamesAt___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__4; @@ -28,6 +29,7 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinitio static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__10; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__12; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__2; +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_packMutual___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -36,30 +38,36 @@ lean_object* lean_mk_array(lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__16; -lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_packMutual___lambda__3___closed__2; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_packMutual___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__8; static lean_object* l_Lean_Elab_WF_withAppN___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__13; +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); extern lean_object* l_instInhabitedNat; lean_object* l_Lean_Meta_ArgsPacker_arities(lean_object*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_WF_packMutual___lambda__1___closed__1; +lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_WF_withAppN(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l_Lean_Meta_ArgsPacker_uncurry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_withAppN___lambda__2___closed__1; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_packMutual___spec__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_WF_withAppN___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, 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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_packMutual___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_getIdx_x3f___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___spec__1___boxed(lean_object*, lean_object*); @@ -75,8 +83,10 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_packMutual___s LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_packMutual___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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*); +extern lean_object* l_Lean_diagnostics; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_WF_packMutual___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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*); @@ -94,6 +104,7 @@ lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_packMutual___lambda__2___closed__1; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); extern lean_object* l_Lean_pp_sanitizeNames; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -102,13 +113,17 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Le LEAN_EXPORT lean_object* l_Lean_Elab_WF_withAppN___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1; static lean_object* l_Lean_Elab_WF_packMutual___lambda__3___closed__4; lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addAsAxiom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3; lean_object* lean_nat_sub(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__1; lean_object* l_Lean_Expr_getAppFn(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_packMutual(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_getIdx_x3f___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___spec__1___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__14; @@ -118,6 +133,8 @@ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_WF_packMutual___lambda__3___closed__3; static lean_object* l_Lean_Elab_WF_packMutual___lambda__3___closed__1; LEAN_EXPORT uint8_t l_Array_getIdx_x3f___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_post___spec__1___lambda__1(lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__4; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_packMutual___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); @@ -139,7 +156,7 @@ static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDef _start: { lean_object* x_1; -x_1 = l_Lean_pp_sanitizeNames; +x_1 = l_Lean_maxRecDepth; return x_1; } } @@ -291,423 +308,602 @@ x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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) { _start: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_12 = l_Lean_Meta_getLevel(x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = 1; +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_10, 4); +lean_dec(x_14); +x_15 = lean_ctor_get(x_10, 2); +lean_dec(x_15); +x_16 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__1; +x_17 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_16); +lean_ctor_set(x_10, 4, x_17); +lean_ctor_set(x_10, 2, x_1); +lean_ctor_set_uint8(x_10, sizeof(void*)*11 + 1, x_2); +lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_16 = l_Lean_Meta_checkpointDefEq___at_Lean_Meta_isLevelDefEq___spec__1(x_1, x_13, x_15, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_18 = lean_infer_type(x_3, x_4, x_5, x_10, x_11, x_12); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_16, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_16); -x_20 = !lean_is_exclusive(x_9); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_9, 2); -x_22 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__1; -x_23 = 0; -x_24 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_21, x_22, x_23); -lean_ctor_set(x_9, 2, x_24); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_25 = lean_infer_type(x_2, x_7, x_8, x_9, x_10, x_19); -if (lean_obj_tag(x_25) == 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_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_5); +lean_inc(x_4); lean_inc(x_6); -x_28 = lean_infer_type(x_6, x_7, x_8, x_9, x_10, x_27); -if (lean_obj_tag(x_28) == 0) +x_21 = lean_infer_type(x_6, x_4, x_5, x_10, x_11, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_ctor_get(x_3, 3); -lean_inc(x_31); -lean_dec(x_3); -x_32 = l_Lean_MessageData_ofName(x_31); -x_33 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__8; -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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__10; -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_indentExpr(x_2); -x_38 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__12; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_MessageData_ofExpr(x_26); -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__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; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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_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_7, 3); +lean_inc(x_24); +lean_dec(x_7); +x_25 = l_Lean_MessageData_ofName(x_24); +x_26 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__8; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__10; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_indentExpr(x_3); +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__12; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = l_Lean_MessageData_ofExpr(x_19); +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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__14; +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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__6; +x_39 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_37); +x_40 = lean_ctor_get(x_8, 3); +lean_inc(x_40); +lean_dec(x_8); +x_41 = l_Lean_MessageData_ofName(x_40); +x_42 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__16; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__6; +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_28); +x_45 = l_Lean_indentExpr(x_6); x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -x_47 = lean_ctor_get(x_4, 3); -lean_inc(x_47); -lean_dec(x_4); -x_48 = l_Lean_MessageData_ofName(x_47); -x_49 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__16; -x_50 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_48); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_32); +x_48 = l_Lean_MessageData_ofExpr(x_22); +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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18; x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_35); -x_52 = l_Lean_indentExpr(x_6); -x_53 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_53, 0, x_51); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_39); -x_55 = l_Lean_MessageData_ofExpr(x_29); -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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18; -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 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_59, 0, x_46); -lean_ctor_set(x_59, 1, x_58); -x_60 = l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(x_59, x_7, x_8, x_9, x_10, x_30); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +x_52 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_52, 0, x_39); +lean_ctor_set(x_52, 1, x_51); +x_53 = l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(x_52, x_4, x_5, x_10, x_11, x_23); +lean_dec(x_11); lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -return x_60; +lean_dec(x_5); +lean_dec(x_4); +return x_53; } else { -uint8_t x_61; -lean_dec(x_26); -lean_dec(x_9); +uint8_t x_54; +lean_dec(x_19); lean_dec(x_10); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_61 = !lean_is_exclusive(x_28); -if (x_61 == 0) +x_54 = !lean_is_exclusive(x_21); +if (x_54 == 0) { -return x_28; +return x_21; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_28, 0); -x_63 = lean_ctor_get(x_28, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_28); -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_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_21, 0); +x_56 = lean_ctor_get(x_21, 1); +lean_inc(x_56); +lean_inc(x_55); +lean_dec(x_21); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } else { -uint8_t x_65; -lean_dec(x_9); +uint8_t x_58; lean_dec(x_10); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_65 = !lean_is_exclusive(x_25); -if (x_65 == 0) +x_58 = !lean_is_exclusive(x_18); +if (x_58 == 0) { -return x_25; +return x_18; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_25, 0); -x_67 = lean_ctor_get(x_25, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_25); -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_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_18, 0); +x_60 = lean_ctor_get(x_18, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_18); +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_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; uint8_t x_81; lean_object* x_82; uint8_t x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_69 = lean_ctor_get(x_9, 0); -x_70 = lean_ctor_get(x_9, 1); -x_71 = lean_ctor_get(x_9, 2); -x_72 = lean_ctor_get(x_9, 3); -x_73 = lean_ctor_get(x_9, 4); -x_74 = lean_ctor_get(x_9, 5); -x_75 = lean_ctor_get(x_9, 6); -x_76 = lean_ctor_get(x_9, 7); -x_77 = lean_ctor_get(x_9, 8); -x_78 = lean_ctor_get(x_9, 9); -x_79 = lean_ctor_get(x_9, 10); -x_80 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -x_81 = lean_ctor_get_uint8(x_9, sizeof(void*)*11 + 1); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_inc(x_72); -lean_inc(x_71); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_62 = lean_ctor_get(x_10, 0); +x_63 = lean_ctor_get(x_10, 1); +x_64 = lean_ctor_get(x_10, 3); +x_65 = lean_ctor_get(x_10, 5); +x_66 = lean_ctor_get(x_10, 6); +x_67 = lean_ctor_get(x_10, 7); +x_68 = lean_ctor_get(x_10, 8); +x_69 = lean_ctor_get(x_10, 9); +x_70 = lean_ctor_get(x_10, 10); +x_71 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); lean_inc(x_70); lean_inc(x_69); -lean_dec(x_9); -x_82 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__1; -x_83 = 0; -x_84 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_71, x_82, x_83); -x_85 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_85, 0, x_69); -lean_ctor_set(x_85, 1, x_70); -lean_ctor_set(x_85, 2, x_84); -lean_ctor_set(x_85, 3, x_72); -lean_ctor_set(x_85, 4, x_73); -lean_ctor_set(x_85, 5, x_74); -lean_ctor_set(x_85, 6, x_75); -lean_ctor_set(x_85, 7, x_76); -lean_ctor_set(x_85, 8, x_77); -lean_ctor_set(x_85, 9, x_78); -lean_ctor_set(x_85, 10, x_79); -lean_ctor_set_uint8(x_85, sizeof(void*)*11, x_80); -lean_ctor_set_uint8(x_85, sizeof(void*)*11 + 1, x_81); -lean_inc(x_10); -lean_inc(x_85); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_2); -x_86 = lean_infer_type(x_2, x_7, x_8, x_85, x_10, x_19); -if (lean_obj_tag(x_86) == 0) -{ -lean_object* x_87; lean_object* x_88; lean_object* x_89; -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); -lean_inc(x_10); -lean_inc(x_85); -lean_inc(x_8); -lean_inc(x_7); +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_dec(x_10); +x_72 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__1; +x_73 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_72); +x_74 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_74, 0, x_62); +lean_ctor_set(x_74, 1, x_63); +lean_ctor_set(x_74, 2, x_1); +lean_ctor_set(x_74, 3, x_64); +lean_ctor_set(x_74, 4, x_73); +lean_ctor_set(x_74, 5, x_65); +lean_ctor_set(x_74, 6, x_66); +lean_ctor_set(x_74, 7, x_67); +lean_ctor_set(x_74, 8, x_68); +lean_ctor_set(x_74, 9, x_69); +lean_ctor_set(x_74, 10, x_70); +lean_ctor_set_uint8(x_74, sizeof(void*)*11, x_71); +lean_ctor_set_uint8(x_74, sizeof(void*)*11 + 1, x_2); +lean_inc(x_11); +lean_inc(x_74); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_75 = lean_infer_type(x_3, x_4, x_5, x_74, x_11, x_12); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +lean_inc(x_11); +lean_inc(x_74); +lean_inc(x_5); +lean_inc(x_4); lean_inc(x_6); -x_89 = lean_infer_type(x_6, x_7, x_8, x_85, x_10, x_88); -if (lean_obj_tag(x_89) == 0) +x_78 = lean_infer_type(x_6, x_4, x_5, x_74, x_11, x_77); +if (lean_obj_tag(x_78) == 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; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_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_ctor_get(x_3, 3); -lean_inc(x_92); -lean_dec(x_3); -x_93 = l_Lean_MessageData_ofName(x_92); -x_94 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__8; -x_95 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_93); -x_96 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__10; -x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_indentExpr(x_2); -x_99 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__12; +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_79 = lean_ctor_get(x_78, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); +x_81 = lean_ctor_get(x_7, 3); +lean_inc(x_81); +lean_dec(x_7); +x_82 = l_Lean_MessageData_ofName(x_81); +x_83 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__8; +x_84 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_84, 0, x_83); +lean_ctor_set(x_84, 1, x_82); +x_85 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__10; +x_86 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_86, 0, x_84); +lean_ctor_set(x_86, 1, x_85); +x_87 = l_Lean_indentExpr(x_3); +x_88 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__12; +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_MessageData_ofExpr(x_76); +x_92 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +x_93 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__14; +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 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__6; +x_96 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +x_97 = lean_ctor_get(x_8, 3); +lean_inc(x_97); +lean_dec(x_8); +x_98 = l_Lean_MessageData_ofName(x_97); +x_99 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__16; +x_100 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_98); x_101 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_101, 0, x_99); -lean_ctor_set(x_101, 1, x_100); -x_102 = l_Lean_MessageData_ofExpr(x_87); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_85); +x_102 = l_Lean_indentExpr(x_6); x_103 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_103, 0, x_101); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__14; -x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -x_106 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__6; -x_107 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_105); -x_108 = lean_ctor_get(x_4, 3); -lean_inc(x_108); +x_104 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_89); +x_105 = l_Lean_MessageData_ofExpr(x_79); +x_106 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_106, 0, x_104); +lean_ctor_set(x_106, 1, x_105); +x_107 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18; +x_108 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_108, 0, x_106); +lean_ctor_set(x_108, 1, x_107); +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_96); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(x_109, x_4, x_5, x_74, x_11, x_80); +lean_dec(x_11); +lean_dec(x_74); +lean_dec(x_5); lean_dec(x_4); -x_109 = l_Lean_MessageData_ofName(x_108); -x_110 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__16; -x_111 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_111, 0, x_110); -lean_ctor_set(x_111, 1, x_109); -x_112 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_96); -x_113 = l_Lean_indentExpr(x_6); -x_114 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_114, 0, x_112); -lean_ctor_set(x_114, 1, x_113); -x_115 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_100); -x_116 = l_Lean_MessageData_ofExpr(x_90); -x_117 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_117, 0, x_115); -lean_ctor_set(x_117, 1, x_116); -x_118 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18; -x_119 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_107); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(x_120, x_7, x_8, x_85, x_10, x_91); -lean_dec(x_10); -lean_dec(x_85); -lean_dec(x_8); -lean_dec(x_7); -return x_121; +return x_110; } else { -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -lean_dec(x_87); -lean_dec(x_85); -lean_dec(x_10); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +lean_dec(x_76); +lean_dec(x_74); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_122 = lean_ctor_get(x_89, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_89, 1); -lean_inc(x_123); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_124 = x_89; +x_111 = lean_ctor_get(x_78, 0); +lean_inc(x_111); +x_112 = lean_ctor_get(x_78, 1); +lean_inc(x_112); +if (lean_is_exclusive(x_78)) { + lean_ctor_release(x_78, 0); + lean_ctor_release(x_78, 1); + x_113 = x_78; } else { - lean_dec_ref(x_89); - x_124 = lean_box(0); + lean_dec_ref(x_78); + x_113 = lean_box(0); } -if (lean_is_scalar(x_124)) { - x_125 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_113)) { + x_114 = lean_alloc_ctor(1, 2, 0); } else { - x_125 = x_124; + x_114 = x_113; } -lean_ctor_set(x_125, 0, x_122); -lean_ctor_set(x_125, 1, x_123); -return x_125; +lean_ctor_set(x_114, 0, x_111); +lean_ctor_set(x_114, 1, x_112); +return x_114; } } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -lean_dec(x_85); -lean_dec(x_10); +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +lean_dec(x_74); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_126 = lean_ctor_get(x_86, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_86, 1); -lean_inc(x_127); -if (lean_is_exclusive(x_86)) { - lean_ctor_release(x_86, 0); - lean_ctor_release(x_86, 1); - x_128 = x_86; +x_115 = lean_ctor_get(x_75, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_75, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_75)) { + lean_ctor_release(x_75, 0); + lean_ctor_release(x_75, 1); + x_117 = x_75; } else { - lean_dec_ref(x_86); - x_128 = lean_box(0); + lean_dec_ref(x_75); + x_117 = lean_box(0); } -if (lean_is_scalar(x_128)) { - x_129 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_117)) { + x_118 = lean_alloc_ctor(1, 2, 0); } else { - x_129 = x_128; + x_118 = x_117; +} +lean_ctor_set(x_118, 0, x_115); +lean_ctor_set(x_118, 1, x_116); +return x_118; +} +} +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_sanitizeNames; +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3() { +_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_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; } -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_127); -return x_129; } +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___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; } } +static lean_object* _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +lean_dec(x_5); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_12 = l_Lean_Meta_getLevel(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; uint8_t x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = 1; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_16 = l_Lean_Meta_checkpointDefEq___at_Lean_Meta_isLevelDefEq___spec__1(x_1, x_13, x_15, x_7, x_8, x_9, x_10, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_unbox(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +lean_object* x_19; 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; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_ctor_get(x_9, 2); +lean_inc(x_20); +x_21 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1; +x_22 = 0; +x_23 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_20, x_21, x_22); +x_24 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2; +x_25 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_23, x_24); +x_26 = lean_st_ref_get(x_10, 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_ctor_get(x_27, 0); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Kernel_isDiagnosticsEnabled(x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +if (x_25 == 0) +{ +x_31 = x_15; +goto block_59; +} else { -uint8_t x_130; +x_31 = x_22; +goto block_59; +} +} +else +{ +if (x_25 == 0) +{ +x_31 = x_22; +goto block_59; +} +else +{ +x_31 = x_15; +goto block_59; +} +} +block_59: +{ +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_st_ref_take(x_10, x_28); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = !lean_is_exclusive(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_36 = lean_ctor_get(x_33, 0); +x_37 = lean_ctor_get(x_33, 4); +lean_dec(x_37); +x_38 = l_Lean_Kernel_enableDiag(x_36, x_25); +x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6; +lean_ctor_set(x_33, 4, x_39); +lean_ctor_set(x_33, 0, x_38); +x_40 = lean_st_ref_set(x_10, x_33, x_34); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_box(0); +x_43 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(x_23, x_25, x_2, x_7, x_8, x_6, x_3, x_4, x_42, x_9, x_10, 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; +x_44 = lean_ctor_get(x_33, 0); +x_45 = lean_ctor_get(x_33, 1); +x_46 = lean_ctor_get(x_33, 2); +x_47 = lean_ctor_get(x_33, 3); +x_48 = lean_ctor_get(x_33, 5); +x_49 = lean_ctor_get(x_33, 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_33); +x_50 = l_Lean_Kernel_enableDiag(x_44, x_25); +x_51 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6; +x_52 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_45); +lean_ctor_set(x_52, 2, x_46); +lean_ctor_set(x_52, 3, x_47); +lean_ctor_set(x_52, 4, x_51); +lean_ctor_set(x_52, 5, x_48); +lean_ctor_set(x_52, 6, x_49); +x_53 = lean_st_ref_set(x_10, x_52, x_34); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = lean_box(0); +x_56 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(x_23, x_25, x_2, x_7, x_8, x_6, x_3, x_4, x_55, x_9, x_10, x_54); +return x_56; +} +} +else +{ +lean_object* x_57; lean_object* x_58; +x_57 = lean_box(0); +x_58 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(x_23, x_25, x_2, x_7, x_8, x_6, x_3, x_4, x_57, x_9, x_10, x_28); +return x_58; +} +} +} +else +{ +uint8_t x_60; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -716,33 +912,33 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_130 = !lean_is_exclusive(x_16); -if (x_130 == 0) -{ -lean_object* x_131; lean_object* x_132; -x_131 = lean_ctor_get(x_16, 0); -lean_dec(x_131); -x_132 = lean_box(0); -lean_ctor_set(x_16, 0, x_132); +x_60 = !lean_is_exclusive(x_16); +if (x_60 == 0) +{ +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_16, 0); +lean_dec(x_61); +x_62 = lean_box(0); +lean_ctor_set(x_16, 0, x_62); return x_16; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_16, 1); -lean_inc(x_133); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_16, 1); +lean_inc(x_63); lean_dec(x_16); -x_134 = lean_box(0); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_134); -lean_ctor_set(x_135, 1, x_133); -return x_135; +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_64); +lean_ctor_set(x_65, 1, x_63); +return x_65; } } } else { -uint8_t x_136; +uint8_t x_66; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -751,29 +947,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_136 = !lean_is_exclusive(x_16); -if (x_136 == 0) +x_66 = !lean_is_exclusive(x_16); +if (x_66 == 0) { return x_16; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_16, 0); -x_138 = lean_ctor_get(x_16, 1); -lean_inc(x_138); -lean_inc(x_137); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_16, 0); +x_68 = lean_ctor_get(x_16, 1); +lean_inc(x_68); +lean_inc(x_67); lean_dec(x_16); -x_139 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_139, 0, x_137); -lean_ctor_set(x_139, 1, x_138); -return x_139; +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; } } } else { -uint8_t x_140; +uint8_t x_70; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -783,23 +979,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_140 = !lean_is_exclusive(x_12); -if (x_140 == 0) +x_70 = !lean_is_exclusive(x_12); +if (x_70 == 0) { return x_12; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_141 = lean_ctor_get(x_12, 0); -x_142 = lean_ctor_get(x_12, 1); -lean_inc(x_142); -lean_inc(x_141); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_12, 0); +x_72 = lean_ctor_get(x_12, 1); +lean_inc(x_72); +lean_inc(x_71); lean_dec(x_12); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -return x_143; +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; } } } @@ -846,7 +1042,7 @@ lean_inc(x_27); lean_inc(x_6); lean_inc(x_7); lean_inc(x_8); -x_28 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___boxed), 11, 4); +x_28 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2), 11, 4); lean_closure_set(x_28, 0, x_8); lean_closure_set(x_28, 1, x_7); lean_closure_set(x_28, 2, x_6); @@ -1161,7 +1357,7 @@ lean_dec(x_18); lean_dec(x_1); x_20 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_20, 0, x_19); -x_21 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_15, x_20, x_16, x_4, x_5, x_6, x_7, x_8); +x_21 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_15, x_20, x_16, x_4, x_5, x_6, x_7, x_8); return x_21; } else @@ -1174,19 +1370,21 @@ lean_dec(x_22); lean_dec(x_1); x_24 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_24, 0, x_23); -x_25 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_15, x_24, x_16, x_4, x_5, x_6, x_7, x_8); +x_25 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_15, x_24, x_16, x_4, x_5, x_6, x_7, x_8); return x_25; } } } } -LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___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) { _start: { -lean_object* x_12; -x_12 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_5); -return x_12; +uint8_t x_13; lean_object* x_14; +x_13 = lean_unbox(x_2); +lean_dec(x_2); +x_14 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1(x_1, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_9); +return x_14; } } LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___boxed(lean_object** _args) { @@ -1410,7 +1608,7 @@ x_23 = lean_alloc_closure((void*)(l_Lean_Elab_WF_withAppN___lambda__2), 10, 3); lean_closure_set(x_23, 0, x_15); lean_closure_set(x_23, 1, x_3); lean_closure_set(x_23, 2, x_18); -x_24 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_20, x_22, x_23, x_4, x_5, x_6, x_7, x_21); +x_24 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_20, x_22, x_23, x_4, x_5, x_6, x_7, x_21); return x_24; } else @@ -3120,7 +3318,7 @@ lean_closure_set(x_19, 2, x_3); lean_closure_set(x_19, 3, x_17); lean_closure_set(x_19, 4, x_11); lean_closure_set(x_19, 5, x_2); -x_20 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_18, x_16, x_19, x_6, x_7, x_8, x_9, x_12); +x_20 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_18, x_16, x_19, x_6, x_7, x_8, x_9, x_12); return x_20; } } @@ -3314,6 +3512,18 @@ l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__L lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__17); l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18(); lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__1___closed__18); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__1); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__2); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__3); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__4 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__4(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__4); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__5); +l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6 = _init_l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6(); +lean_mark_persistent(l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_PackMutual_0__Lean_Elab_WF_checkCodomainsLevel___spec__1___lambda__2___closed__6); l_Lean_Elab_WF_withAppN___lambda__2___closed__1 = _init_l_Lean_Elab_WF_withAppN___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_WF_withAppN___lambda__2___closed__1); l_Lean_Elab_WF_withAppN___lambda__2___closed__2 = _init_l_Lean_Elab_WF_withAppN___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index 38a2b8de83b0..9d96d1e34f40 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -17,7 +17,6 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxCo LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__5(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -29,23 +28,22 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFie LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___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*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_toVisibility___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___lambda__1(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___closed__2; lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__8; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instDecidableEqStructFieldKind___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_toCtorIdx___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__4; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); @@ -73,18 +71,19 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru lean_object* l_Lean_Elab_Command_checkResultingUniverse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_getStuckMVar_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26; lean_object* l_Lean_Elab_Command_runTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_sortDeclLevelParams(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___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*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__11___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_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6; extern lean_object* l_Lean_declRangeExt; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__5___closed__1; lean_object* l_Lean_Elab_Command_mkResultUniverse(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___spec__1(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__1; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__14; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___closed__6; @@ -96,12 +95,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__3 static lean_object* l_Lean_Elab_Command_elabStructure___closed__11; extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_updateFieldInfoVal___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields___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_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__4; lean_object* l_Lean_indentD(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__10; @@ -110,11 +108,9 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure lean_object* l_Lean_Elab_Command_getMainModule___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___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_Structure___hyg_13358____closed__9; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___closed__4; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_setReducibilityStatusCore(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -124,17 +120,18 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure lean_object* l_Lean_Level_addOffsetAux(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInStructure___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAutoParamTactic_x3f(lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsed___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___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_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_findDocString_x3f(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__1; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__6(lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___closed__1; lean_object* l_Lean_Syntax_getId(lean_object*); lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -142,7 +139,6 @@ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withUsed___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_Expr_sort___override(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___closed__3; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__7; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -150,14 +146,15 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_ 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___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSourceInstImplicit(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__13; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__5; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__8; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__5; lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); uint8_t l_Array_contains___at_Lean_findField_x3f___spec__1(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__3; lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f___closed__1; lean_object* lean_mk_array(lean_object*, lean_object*); @@ -170,8 +167,8 @@ lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instReprStructFieldInfo___closed__1; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____boxed(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_StructFieldInfo_isFromParent(lean_object*); uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___spec__1(lean_object*, lean_object*); @@ -180,24 +177,23 @@ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Expr_bvar___override(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInStructure___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_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22; extern lean_object* l_Lean_Expr_instBEq; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19; static lean_object* l_Lean_Elab_Command_elabStructure___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___closed__1; lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_containsFieldName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___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_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__9; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCtor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___closed__1; static lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__4___closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__6; lean_object* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_structureDiamondWarning; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -212,9 +208,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_el lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed___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_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15; lean_object* l_Lean_mkIdentFrom(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542_(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_updateFieldInfoVal___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__5; @@ -222,13 +221,12 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getResu static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_updateResultingUniverse___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___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* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedStructFieldInfo; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed___closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__8; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParamFVars___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*); @@ -238,10 +236,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkSynthetic(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21; static uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__2___boxed(lean_object**); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__8; @@ -250,14 +248,15 @@ static lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Command_elab static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2; static uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_utf8_byte_size(lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_findExistingField_x3f___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_CollectLevelParams_main(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__9(lean_object*); @@ -266,10 +265,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at___private_Lean_Ela LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___boxed__const__1; lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_GeneralizeTelescope_generalizeTelescopeAux___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__14; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__3___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_Structure___hyg_3469____closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findExistingField_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkDefaultFnOfProjFn(lean_object*); @@ -281,14 +281,17 @@ lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__9___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_instDecidableEqStructFieldKind(uint8_t, uint8_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5; lean_object* l_Lean_getDefaultFnForField_x3f(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTermEnsuringType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__3; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__2___closed__1; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__9; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3; static lean_object* l_Lean_getDocStringText___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__14___closed__3; 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*); size_t lean_ptr_addr(lean_object*); @@ -297,13 +300,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_fi static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__2___closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__9___rarg___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__2; size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__3; uint8_t l_Lean_TagDeclarationExtension_isTagged(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_removeUnused___closed__1; @@ -312,9 +315,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_le lean_object* lean_expr_abstract(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__7___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_Structure_0__Lean_Elab_Command_elabStructureView___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_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__4(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*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -338,10 +343,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__3(lean_obje uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_panic___at_Lean_Meta_whnfCore_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__2; +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____boxed(lean_object*, lean_object*); lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_Term_elabLetDeclAux___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getDeclarationSelectionRef(lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6; uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam___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*); @@ -376,11 +382,11 @@ lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_obje LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17; 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*); lean_object* l_Lean_registerStructure(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f___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_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f___boxed(lean_object*, 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_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -388,7 +394,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_toCtorIdx(uint8_t); lean_object* l_Lean_LocalContext_setBinderInfo(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_getRef(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__2___closed__4; @@ -397,17 +402,17 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Struc static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__7; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__3___closed__1; lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23; lean_object* l_Lean_Elab_toAttributeKind___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13; lean_object* l_Lean_Elab_Term_applyAttributesAt(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___closed__4; lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getBetterRef(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21; lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Term_expandDeclId___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_Structure_0__Lean_Elab_Command_elabStructureView___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -416,6 +421,7 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___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_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_defaultCtorName___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addCtorFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_value(lean_object*); @@ -423,13 +429,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5___closed__3; lean_object* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getLevel(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22; lean_object* l_Lean_Meta_reduceProjOf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___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_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6; static lean_object* l_Lean_Elab_Command_elabStructure___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor(lean_object*, lean_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_initFn____x40_Lean_Elab_Structure___hyg_13361_(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__11; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__8; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -438,10 +445,10 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0_ static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__4___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1___lambda__1___closed__1; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12; LEAN_EXPORT uint8_t l___private_Lean_Elab_Structure_0__Lean_Elab_Command_validStructType(lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___closed__2; lean_object* l_Lean_Expr_inferImplicit(lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4; lean_object* l_Lean_Meta_mkAppM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_removeUnused(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -449,8 +456,10 @@ lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageCo LEAN_EXPORT uint8_t l_Lean_Elab_Command_StructFieldKind_ofNat(lean_object*); lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__6; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___boxed(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__5; lean_object* l_Lean_Syntax_getOptional_x3f(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -464,6 +473,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mk LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__3; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -478,6 +488,7 @@ static lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_El static lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1___closed__1; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__10(lean_object*, size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__6; @@ -486,16 +497,15 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Structure_ lean_object* l_Lean_Meta_mkId(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_DerivingClassView_applyHandlers(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__2; static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__3___closed__1; extern lean_object* l_Lean_levelZero; lean_object* l_Lean_Elab_Command_shouldInferResultUniverse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8; lean_object* l_Lean_Elab_expandDeclSig(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___lambda__1___boxed(lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__4; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4; lean_object* l_Lean_Meta_mkInjectiveTheorems(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__5___closed__2; @@ -505,23 +515,23 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structur 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*); lean_object* l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16; lean_object* l_Lean_Elab_Command_getCurrMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___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_Structure_0__Lean_Elab_Command_addDefaults___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__11; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addCtorFields___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25; lean_object* l_panic___at_Lean_Expr_appFn_x21___spec__1(lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getSepArgs(lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___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*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier___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_Structure___hyg_3469____closed__3; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__4; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__6; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); @@ -531,17 +541,19 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFiel lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__5(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358_(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectLevelParamsInStructure___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23; lean_object* l_Lean_Elab_Term_getLevelNames___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instReprStructFieldKind; static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_addDeclarationRanges___at_Lean_Elab_Command_elabStructure___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsed___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_throwErrorAt___at_Lean_Elab_Term_processDefDeriving___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___closed__1; @@ -552,7 +564,6 @@ uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler static lean_object* l_Lean_Elab_Command_elabStructure___closed__8; static lean_object* l_Lean_getDocStringText___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__14___closed__1; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__19; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__10; @@ -568,13 +579,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___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_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*); static lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1___lambda__1___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___lambda__1___closed__1; lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_containsFieldName___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___closed__1; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__1; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___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_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7; +lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__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_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -588,7 +603,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_wi LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___closed__2; static lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findExistingField_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_mkNoConfusionEnum_mkToCtorIdx___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -602,6 +616,7 @@ static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Lean_Elab_Command_instReprStructFieldInfo; lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4; 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); @@ -610,13 +625,13 @@ lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4(lea LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); 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*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469_(lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__1(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_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__2(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_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334_(uint8_t, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__2; static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__4; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__2; @@ -630,8 +645,8 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11; LEAN_EXPORT lean_object* l_Array_filterMapM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___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___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParamFVar(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_elabStructure___closed__9; @@ -643,9 +658,11 @@ lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___closed__3; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkCasesOn___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Elab_Modifiers_isPrivate(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -662,13 +679,12 @@ lean_object* l_Lean_FileMap_leanPosToLspPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_failed(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_Structure_0__Lean_Elab_Command_elabFieldTypeValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getResultUniverse(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_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__6; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasAnyFVar_visit___at_Lean_Expr_containsFVar___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_defaultCtorName; @@ -688,7 +704,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_ad LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getOptDerivingClasses(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539_(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidInductiveModifier___at_Lean_Elab_Command_elabStructure___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__1; @@ -698,11 +713,9 @@ LEAN_EXPORT uint8_t l_Lean_Elab_Command_StructFieldInfo_isSubobject(lean_object* extern lean_object* l_Lean_NameSet_empty; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__10; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAtomic(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParam_x27___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -715,22 +728,18 @@ lean_object* l_Option_repr___at___private_Lean_Structure_0__Lean_reprStructureFi lean_object* l_Lean_Expr_getForallBody(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkFreshLevelMVarsFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__3___boxed(lean_object**); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__1; lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__14; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_processOveriddenDefaultValues___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__3; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkToParentName___closed__1; lean_object* l_Lean_Name_beq___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabFieldTypeValue___lambda__1___boxed(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instantiateLevelMVars___at_Lean_Elab_Command_accLevelAtCtor___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -743,6 +752,7 @@ lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_obje static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__3; static lean_object* l_Lean_Elab_elabAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__3___closed__1; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -753,13 +763,13 @@ lean_object* l_Lean_Elab_Term_elabBinders___rarg(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_findExistingField_x3f___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_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getNestedProjectionArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParamFVar___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_Structure_0__Lean_Elab_Command_reduceProjs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___closed__2; uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_updateResultingUniverse(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___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_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__8; @@ -778,14 +788,14 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_setSour LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents_go(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__6(lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyDefaultValue_x3f_go_x3f___closed__2; lean_object* l_Lean_ofExceptKernelException___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkAuxConstructions___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_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___closed__5; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__6; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_levelMVarToParam_levelMVarToParamFVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_updateResultingUniverse___closed__1; @@ -798,14 +808,13 @@ static lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure static lean_object* l_Lean_Elab_Command_instInhabitedStructFieldInfo___closed__1; extern lean_object* l_Lean_Expr_instHashable; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___closed__4; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_removeUnused___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__13; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Command_elabStructure___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4; 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_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__5___boxed(lean_object**); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337_(uint8_t, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabDeclAttrs___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -821,7 +830,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Struc lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isPrivateName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_StructFieldKind_noConfusion(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__4; LEAN_EXPORT lean_object* l_Lean_setEnv___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections___spec__1(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_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__3; @@ -844,7 +852,6 @@ extern lean_object* l_Lean_protectedExt; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__3; lean_object* lean_erase_macro_scopes(lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__9; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__2___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___lambda__1___boxed(lean_object*); @@ -853,17 +860,16 @@ static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy(lean_object*); lean_object* l_Lean_Meta_mkAuxDefinition(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getResultUniverse___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withParents___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_Command_StructFieldKind_ofNat___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_removeUnused(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__12; lean_object* l_Lean_getStructureCtor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_toVisibility(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___closed__1; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7; uint8_t l_Lean_Level_isMVar(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__1; static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__11___closed__1; @@ -872,17 +878,15 @@ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFie static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__5; static lean_object* l_Lean_Elab_Command_checkValidFieldModifier___lambda__2___closed__2; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__5; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withUsed(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___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_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__2(lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkToParentName(lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___closed__6; @@ -891,7 +895,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__7(lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___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*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_getDeclarationRange___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__19(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getNestedProjectionArg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -905,9 +908,7 @@ static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structu LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__3; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__1; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__5___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_Structure___hyg_3466____closed__4; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___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_string_append(lean_object*, lean_object*); @@ -916,17 +917,15 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkAuxConstructions___closed__1; lean_object* l_Lean_mkAtomFrom(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25; lean_object* l___private_Lean_Expr_0__Lean_reprExpr____x40_Lean_Expr___hyg_2961_(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instReprStructFieldKind___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13; static lean_object* l_Lean_Elab_Command_elabStructure___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__1; static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__5; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__3___closed__2; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__3; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -944,18 +943,17 @@ static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__L LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsed(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_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__20; static lean_object* l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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*); lean_object* l_Lean_Meta_mkProjection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__4; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUsed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__5; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_mkCompositeField___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466_(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__7(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields(lean_object*); @@ -969,11 +967,12 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__12(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_Structure_0__Lean_Elab_Command_expandCtor___spec__4___closed__3; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25; uint8_t l_Lean_Expr_isFVar(lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__4___boxed(lean_object**); lean_object* l_Lean_expandMacros(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__12; lean_object* l_Lean_getStructureFieldsFlattened(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___lambda__5(lean_object*, 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*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___closed__10; @@ -991,12 +990,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mk static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__3; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent___closed__3; -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__1___lambda__1___boxed(lean_object**); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__9; uint8_t l_Lean_Elab_Modifiers_isPartial(lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__7; static lean_object* l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1005,17 +1003,19 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Stru static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__4___closed__1; lean_object* l_Lean_Expr_collectFVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___spec__2___closed__2; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reduceProjs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___lambda__1___boxed(lean_object**); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_addAuxDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__18(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10; lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandFields___spec__5___lambda__6___closed__3; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_processSubfields_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_addDefaults___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1; +static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___boxed(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_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent_copyFields___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*); @@ -1207,7 +1207,7 @@ x_6 = lean_box(x_5); return x_6; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__1() { _start: { lean_object* x_1; @@ -1215,17 +1215,17 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.Command.StructFieldKind.newField", 42 return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -1234,23 +1234,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1258,7 +1258,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -1267,23 +1267,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1291,7 +1291,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9() { _start: { lean_object* x_1; @@ -1299,33 +1299,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.Command.StructFieldKind.copiedField", return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1333,23 +1333,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1357,7 +1357,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15() { _start: { lean_object* x_1; @@ -1365,33 +1365,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.Command.StructFieldKind.fromParent", return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1399,23 +1399,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1423,7 +1423,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21() { _start: { lean_object* x_1; @@ -1431,33 +1431,33 @@ x_1 = lean_mk_string_from_bytes("Lean.Elab.Command.StructFieldKind.subobject", 4 return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1465,23 +1465,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -1489,7 +1489,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -1501,14 +1501,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5; +x_5 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8; +x_7 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -1521,14 +1521,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12; +x_11 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14; +x_13 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -1541,14 +1541,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18; +x_17 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20; +x_19 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -1561,14 +1561,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24; +x_23 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26; +x_25 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -1576,13 +1576,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334_(x_3, x_2); +x_4 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -1591,7 +1591,7 @@ static lean_object* _init_l_Lean_Elab_Command_instReprStructFieldKind___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____boxed), 2, 0); return x_1; } } @@ -1645,7 +1645,7 @@ x_1 = l_Lean_Elab_Command_instInhabitedStructFieldInfo___closed__2; return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__1() { _start: { lean_object* x_1; @@ -1653,29 +1653,29 @@ x_1 = lean_mk_string_from_bytes("name", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__4() { _start: { lean_object* x_1; @@ -1683,29 +1683,29 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3; -x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3; +x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____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; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -1714,7 +1714,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8() { _start: { lean_object* x_1; @@ -1722,17 +1722,17 @@ x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10() { _start: { lean_object* x_1; @@ -1740,17 +1740,17 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -1759,7 +1759,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13() { _start: { lean_object* x_1; @@ -1767,17 +1767,17 @@ x_1 = lean_mk_string_from_bytes("fvar", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15() { _start: { lean_object* x_1; @@ -1785,17 +1785,17 @@ x_1 = lean_mk_string_from_bytes("kind", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17() { _start: { lean_object* x_1; @@ -1803,17 +1803,17 @@ x_1 = lean_mk_string_from_bytes("value\?", 6); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19() { _start: { lean_object* x_1; lean_object* x_2; @@ -1822,7 +1822,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20() { _start: { lean_object* x_1; @@ -1830,35 +1830,35 @@ x_1 = lean_mk_string_from_bytes("{ ", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24() { _start: { lean_object* x_1; @@ -1866,17 +1866,17 @@ x_1 = lean_mk_string_from_bytes(" }", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25() { +static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24; +x_1 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; @@ -1884,7 +1884,7 @@ x_3 = lean_ctor_get(x_1, 0); lean_inc(x_3); x_4 = lean_unsigned_to_nat(0u); x_5 = l_Lean_Name_reprPrec(x_3, x_4); -x_6 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7; +x_6 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -1892,11 +1892,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6; +x_10 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9; +x_12 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -1904,18 +1904,18 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11; +x_16 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11; 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_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5; +x_18 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5; 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 = lean_ctor_get(x_1, 1); lean_inc(x_20); x_21 = l_Lean_Name_reprPrec(x_20, x_4); -x_22 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12; +x_22 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12; x_23 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); @@ -1931,7 +1931,7 @@ lean_ctor_set(x_26, 1, x_12); x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_14); -x_28 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14; +x_28 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); @@ -1956,7 +1956,7 @@ lean_ctor_set(x_36, 1, x_12); x_37 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_37, 0, x_36); lean_ctor_set(x_37, 1, x_14); -x_38 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16; +x_38 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16; x_39 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); @@ -1964,7 +1964,7 @@ x_40 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_40, 0, x_39); lean_ctor_set(x_40, 1, x_18); x_41 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -x_42 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334_(x_41, x_4); +x_42 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337_(x_41, x_4); x_43 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_43, 0, x_6); lean_ctor_set(x_43, 1, x_42); @@ -1980,7 +1980,7 @@ lean_ctor_set(x_46, 1, x_12); x_47 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_47, 0, x_46); lean_ctor_set(x_47, 1, x_14); -x_48 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18; +x_48 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18; x_49 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_49, 0, x_47); lean_ctor_set(x_49, 1, x_48); @@ -1991,7 +1991,7 @@ x_51 = lean_ctor_get(x_1, 3); lean_inc(x_51); lean_dec(x_1); x_52 = l_Option_repr___at___private_Lean_Structure_0__Lean_reprStructureFieldInfo____x40_Lean_Structure___hyg_62____spec__2(x_51, x_4); -x_53 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19; +x_53 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19; x_54 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -2001,15 +2001,15 @@ lean_ctor_set_uint8(x_55, sizeof(void*)*1, x_8); x_56 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_56, 0, x_50); lean_ctor_set(x_56, 1, x_55); -x_57 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23; +x_57 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23; x_58 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_58, 0, x_57); lean_ctor_set(x_58, 1, x_56); -x_59 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25; +x_59 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25; 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 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22; +x_61 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22; x_62 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_62, 0, x_61); lean_ctor_set(x_62, 1, x_60); @@ -2019,11 +2019,11 @@ lean_ctor_set_uint8(x_63, sizeof(void*)*1, x_8); return x_63; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539_(x_1, x_2); +x_3 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -2032,7 +2032,7 @@ static lean_object* _init_l_Lean_Elab_Command_instReprStructFieldInfo___closed__ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____boxed), 2, 0); return x_1; } } @@ -7625,7 +7625,7 @@ x_68 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_68, 0, x_58); lean_ctor_set(x_68, 1, x_66); lean_ctor_set(x_68, 2, x_64); -x_69 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8; +x_69 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8; lean_inc(x_58); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_58); @@ -7757,7 +7757,7 @@ x_112 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_112, 0, x_102); lean_ctor_set(x_112, 1, x_110); lean_ctor_set(x_112, 2, x_108); -x_113 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8; +x_113 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8; lean_inc(x_102); x_114 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_114, 0, x_102); @@ -7974,7 +7974,7 @@ x_177 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_177, 0, x_167); lean_ctor_set(x_177, 1, x_175); lean_ctor_set(x_177, 2, x_173); -x_178 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8; +x_178 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8; lean_inc(x_167); x_179 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_179, 0, x_167); @@ -9637,7 +9637,7 @@ lean_dec(x_1); return x_8; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__1() { _start: { lean_object* x_1; @@ -9645,17 +9645,17 @@ x_1 = lean_mk_string_from_bytes("structureDiamondWarning", 23); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____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_Structure___hyg_3466____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__3() { _start: { lean_object* x_1; @@ -9663,13 +9663,13 @@ x_1 = lean_mk_string_from_bytes("enable/disable warning messages for structure d return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____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_getDocStringText___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__14___closed__3; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__3; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -9678,7 +9678,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5() { _start: { lean_object* x_1; @@ -9686,25 +9686,25 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5; x_3 = l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__1; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__4; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__4; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6; x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } @@ -23506,7 +23506,7 @@ x_15 = lean_ctor_get(x_11, 4); lean_dec(x_15); x_16 = 0; x_17 = lean_box(0); -x_18 = l_Lean_setReducibilityStatusCore(x_14, x_1, x_2, x_16, x_17); +x_18 = l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(x_14, x_1, x_2, x_16, x_17); x_19 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__5; lean_ctor_set(x_11, 4, x_19); lean_ctor_set(x_11, 0, x_18); @@ -23611,7 +23611,7 @@ lean_inc(x_46); lean_dec(x_11); x_52 = 0; x_53 = lean_box(0); -x_54 = l_Lean_setReducibilityStatusCore(x_46, x_1, x_2, x_52, x_53); +x_54 = l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(x_46, x_1, x_2, x_52, x_53); x_55 = l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__5; x_56 = lean_alloc_ctor(0, 7, 0); lean_ctor_set(x_56, 0, x_54); @@ -27922,7 +27922,7 @@ static lean_object* _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_e _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5; x_2 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -30392,7 +30392,7 @@ lean_dec(x_8); return x_15; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -30402,27 +30402,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_Structure___hyg_13358____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____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_Structure___hyg_13358____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__2; x_2 = l_Lean_Elab_elabModifiers___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__1___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4() { _start: { lean_object* x_1; @@ -30430,17 +30430,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_Structure___hyg_13358____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4; 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_Structure___hyg_13358____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__6() { _start: { lean_object* x_1; @@ -30448,37 +30448,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__7; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__7; x_2 = l_Lean_Elab_elabAttr___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__4___lambda__1___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__8; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__8; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____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_Structure___hyg_13358____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10() { _start: { lean_object* x_1; @@ -30486,17 +30486,17 @@ x_1 = lean_mk_string_from_bytes("Structure", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__9; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__10; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__9; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__12() { _start: { lean_object* x_1; @@ -30504,33 +30504,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_Structure___hyg_13358____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__11; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__11; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____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_Structure___hyg_13358____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__13; -x_2 = lean_unsigned_to_nat(13358u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__13; +x_2 = lean_unsigned_to_nat(13361u); 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_Structure___hyg_13358_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361_(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_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__14; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -30595,58 +30595,58 @@ lean_dec_ref(res); l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___closed__1 = _init_l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_StructFieldKind_noConfusion___rarg___closed__1); l_Lean_Elab_Command_instInhabitedStructFieldKind = _init_l_Lean_Elab_Command_instInhabitedStructFieldKind(); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__1); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__2); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__3); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__4); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__5); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__6); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__7); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__8); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__9); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__10); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__11); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__12); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__13); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__14); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__15); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__16); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__17); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__18); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__19); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__20); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__21); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__22); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__23); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__24); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__25); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_334____closed__26); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__1); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__2); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__3); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__4); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__5); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__6); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__7); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__8); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__9); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__10); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__11); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__12); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__13); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__14); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__15); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__16); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__17); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__18); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__19); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__20); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__21); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__22); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__24); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__25); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26); l_Lean_Elab_Command_instReprStructFieldKind___closed__1 = _init_l_Lean_Elab_Command_instReprStructFieldKind___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instReprStructFieldKind___closed__1); l_Lean_Elab_Command_instReprStructFieldKind = _init_l_Lean_Elab_Command_instReprStructFieldKind(); @@ -30659,56 +30659,56 @@ l_Lean_Elab_Command_instInhabitedStructFieldInfo___closed__2 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Command_instInhabitedStructFieldInfo___closed__2); l_Lean_Elab_Command_instInhabitedStructFieldInfo = _init_l_Lean_Elab_Command_instInhabitedStructFieldInfo(); lean_mark_persistent(l_Lean_Elab_Command_instInhabitedStructFieldInfo); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__1); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__2); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__3); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__4); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__5); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__6); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__7); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__8); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__9); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__10); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__11); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__12); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__13); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__14); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__15); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__16); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__17); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__18); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__19); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__20); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__21); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__22); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__23); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__24); -l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25(); -lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_539____closed__25); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__1 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__1(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__1); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__2 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__2(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__2); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__3); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__4 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__4(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__4); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__5); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__6); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__7); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__8); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__9); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__10); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__11); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__12); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__13); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__14); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__15); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__16); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__17); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__18); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__19); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__20); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__21); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__22); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__23); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__24); +l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25(); +lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldInfo____x40_Lean_Elab_Structure___hyg_542____closed__25); l_Lean_Elab_Command_instReprStructFieldInfo___closed__1 = _init_l_Lean_Elab_Command_instReprStructFieldInfo___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_instReprStructFieldInfo___closed__1); l_Lean_Elab_Command_instReprStructFieldInfo = _init_l_Lean_Elab_Command_instReprStructFieldInfo(); @@ -30954,19 +30954,19 @@ l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed_ lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__2); l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__3 = _init_l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__3(); lean_mark_persistent(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_findFieldInfo_x3f___closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466____closed__6); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3466_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469____closed__6); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_3469_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_structureDiamondWarning = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_structureDiamondWarning); @@ -31241,35 +31241,35 @@ l_Lean_Elab_Command_elabStructure___closed__11 = _init_l_Lean_Elab_Command_elabS lean_mark_persistent(l_Lean_Elab_Command_elabStructure___closed__11); l_Lean_Elab_Command_elabStructure___closed__12 = _init_l_Lean_Elab_Command_elabStructure___closed__12(); lean_mark_persistent(l_Lean_Elab_Command_elabStructure___closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358____closed__14); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13358_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361____closed__14); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Structure___hyg_13361_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Syntax.c b/stage0/stdlib/Lean/Elab/Syntax.c index 7ea307f7611a..70ae37e3a051 100644 --- a/stage0/stdlib/Lean/Elab/Syntax.c +++ b/stage0/stdlib/Lean/Elab/Syntax.c @@ -51,7 +51,6 @@ 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; @@ -564,6 +563,7 @@ static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Elab_Syntax_ static lean_object* l_Lean_Elab_Term_toParserDescr_processNonReserved___closed__11; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_toParserDescr_processAlias___spec__6___lambda__1___closed__11; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_ensureUnaryOutput___closed__21; static lean_object* l_Lean_Elab_Term_toParserDescr_processNullaryOrCat___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Array_sequenceMap_loop___at_Lean_Elab_Command_elabSyntax___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -15064,7 +15064,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__12(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__11(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 88c276b44fe6..ffe2bb45fe0e 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Basic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Basic.c @@ -15,7 +15,6 @@ extern "C" { #endif 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_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -23,32 +22,31 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_ lean_object* l_Lean_Elab_Term_logUnassignedUsingErrorInfos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2___closed__2; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1___closed__1; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__9; 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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__9; static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___closed__1; extern lean_object* l_Lean_profiler; uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; lean_object* lean_mk_empty_array_with_capacity(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__9; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___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_EXPORT lean_object* l_Lean_Elab_admitGoal___lambda__1(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_5514____closed__1; 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_object* l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_TacticM_runCore_x27(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, 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__4(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*, 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*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___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*); 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; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkInitialTacticInfo___elambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain(lean_object*); lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -67,10 +65,10 @@ static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrA LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTactic___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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* 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*); @@ -78,7 +76,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx(lean_object*, le static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___boxed(lean_object**); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,7 +85,6 @@ static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__10; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_evalTactic_expandEval___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_Core_getMessageLog___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__16; LEAN_EXPORT lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Tactic_closeMainGoal___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___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*); @@ -98,8 +94,6 @@ 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*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); @@ -113,17 +107,19 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Tactic_e static lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__3; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_focus(lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaFinishingTactic___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_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___closed__1; +static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__4; static lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___closed__2; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_throwNoGoalsToBeSolved___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__2; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainTag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic_x3f(lean_object*); @@ -140,6 +136,7 @@ uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_ lean_object* l_Lean_MetavarContext_setMVarUserName(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__6; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__16; lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -149,18 +146,20 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic(lean_object*); lean_object* l_Lean_Exception_toMessageData(lean_object*); static lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__12; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_run(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_throwNoGoalsToBeSolved___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainModule___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkInitialTacticInfo___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5559_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadTacticM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withCaseRef___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveState___boxed(lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_throwExs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1(lean_object*, 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__6; @@ -194,23 +193,26 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__1(le LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_expandMacroImpl_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic_throwExs___closed__2; -lean_object* l_Lean_log___at_Lean_Meta_reportDiag___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__2; lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaMAtMain___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___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_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__9___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_Term_reportUnsolvedGoals___closed__2; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__5; lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tryTactic_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM___closed__1; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getGoals(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getGoals___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__5; 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_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*); @@ -227,7 +229,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux___rarg(lean_object LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_adaptExpander(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadTacticM; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__5(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*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMainContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instMonadBacktrackSavedStateTacticM; @@ -236,19 +237,22 @@ LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tacti LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_withMacroExpansion___spec__1(lean_object*); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__3; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__8; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getMainTarget(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__1; uint8_t l_List_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg(lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__1; lean_object* lean_st_mk_ref(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__10; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_withMainContext___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); @@ -257,17 +261,17 @@ 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_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_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__12; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__14; static lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5559____closed__1; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__9(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Term_reportUnsolvedGoals___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,10 +286,11 @@ static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__1; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__3; +extern lean_object* l_Lean_trace_profiler_threshold; uint8_t l_Lean_MetavarContext_isAnonymousMVar(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_ensureHasNoMVars___closed__1; static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_liftMetaTacticAux(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); @@ -293,7 +298,6 @@ 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*); @@ -307,13 +311,11 @@ static lean_object* l_Lean_Elab_Tactic_evalTactic___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___rarg___boxed(lean_object*, lean_object*); uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__6; 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514_(lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__13; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__4; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__6___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_pruneSolvedGoals___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -322,35 +324,37 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveState___rarg(lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getNameOfIdent_x27___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeUsingOrAdmit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__15; -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_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_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*); 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*); -static double l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5556____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkTacticInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_evalTactic___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___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_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_getCurrMacroScope___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___closed__1; 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*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Tactic_evalTactic___spec__9___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_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__6; 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_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517_(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_Elab_Tactic_throwNoGoalsToBeSolved___rarg___closed__2; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Elab_Tactic_evalTactic___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__2; static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__2; lean_object* l_Lean_PersistentHashMap_insert___at_Lean_MVarId_assign___spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_expandEval___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instFunctorOfMonad___rarg(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); @@ -359,6 +363,7 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Tactic_evalTactic_expandEval lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic_throwExs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Tactic_evalTactic_handleEx___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Tactic_closeUsingOrAdmit___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMacroExpansion(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_evalTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -366,6 +371,7 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoal LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Tactic_withMacroExpansion___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__14; 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_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*); @@ -380,8 +386,11 @@ static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_tagUntaggedGoals___spec__2___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_withMainContext(lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTactic___at_Lean_Elab_Tactic_done___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_instAlternativeTacticM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__7; static lean_object* l_Lean_Elab_Tactic_instOrElseTacticM___closed__1; static lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___closed__2; lean_object* l_List_reverse___rarg(lean_object*); @@ -396,12 +405,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__1(lean_object*, size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_getMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic_throwExs___closed__1; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5556_(lean_object*); static lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM___closed__2; static lean_object* l_Lean_Elab_Tactic_instMonadTacticM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instAlternativeTacticM; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_closeMainGoal___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_done___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -430,19 +437,17 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Basic_0__Lean_Elab_Tactic_ LEAN_EXPORT 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*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Tactic_evalTactic_expandEval___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_Term_reportUnsolvedGoals___closed__1; -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__4; static lean_object* l_Lean_Elab_Tactic_mkTacticAttribute___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_instMonadExceptExceptionTacticM___closed__2; static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__4___closed__4; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTactic_handleEx___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__11; 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; @@ -452,16 +457,15 @@ 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*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__7; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__5; lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_saveState(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalTactic___lambda__5___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Tactic_withTacticInfoContext___spec__2___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5514____closed__3; +static lean_object* l_Lean_Elab_Tactic_initFn____x40_Lean_Elab_Tactic_Basic___hyg_5517____closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_run___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT 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*); @@ -822,7 +826,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_15 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_13, x_14, x_2, x_3, x_4, x_5, x_6); +x_15 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_13, x_14, x_2, x_3, x_4, x_5, x_6); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); @@ -6764,119 +6768,7 @@ x_8 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_8; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_io_mono_nanos_now(x_10); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = lean_io_mono_nanos_now(x_16); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) -{ -lean_object* x_19; uint8_t x_20; lean_object* x_21; double x_22; double x_23; double x_24; double x_25; double x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_19 = lean_ctor_get(x_17, 0); -x_20 = 0; -x_21 = lean_unsigned_to_nat(0u); -x_22 = l_Float_ofScientific(x_12, x_20, x_21); -lean_dec(x_12); -x_23 = l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; -x_24 = lean_float_div(x_22, x_23); -x_25 = l_Float_ofScientific(x_19, x_20, x_21); -lean_dec(x_19); -x_26 = lean_float_div(x_25, x_23); -x_27 = lean_box_float(x_24); -x_28 = lean_box_float(x_26); -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_15); -lean_ctor_set(x_30, 1, x_29); -lean_ctor_set(x_17, 0, x_30); -return x_17; -} -else -{ -lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; double x_35; double x_36; double x_37; double x_38; double 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_17, 0); -x_32 = lean_ctor_get(x_17, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = 0; -x_34 = lean_unsigned_to_nat(0u); -x_35 = l_Float_ofScientific(x_12, x_33, x_34); -lean_dec(x_12); -x_36 = l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_evalTactic___spec__5___closed__1; -x_37 = lean_float_div(x_35, x_36); -x_38 = l_Float_ofScientific(x_31, x_33, x_34); -lean_dec(x_31); -x_39 = lean_float_div(x_38, x_36); -x_40 = lean_box_float(x_37); -x_41 = lean_box_float(x_39); -x_42 = lean_alloc_ctor(0, 2, 0); -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_15); -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_32); -return x_44; -} -} -else -{ -uint8_t x_45; -lean_dec(x_12); -x_45 = !lean_is_exclusive(x_14); -if (x_45 == 0) -{ -return x_14; -} -else -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_14, 0); -x_47 = lean_ctor_get(x_14, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_14); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} -} -} -} -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, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; @@ -7145,7 +7037,7 @@ return x_103; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_1) == 0) @@ -7170,89 +7062,23 @@ return x_14; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; -x_11 = lean_apply_9(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_11, 0, x_14); -return x_11; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_11, 0); -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_11); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -} -else -{ -uint8_t x_19; -x_19 = !lean_is_exclusive(x_11); -if (x_19 == 0) -{ -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_11, 0); -x_21 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set_tag(x_11, 0); -lean_ctor_set(x_11, 0, x_21); -return x_11; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_11, 0); -x_23 = lean_ctor_get(x_11, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_11); -x_24 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_24, 0, x_22); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -return x_25; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_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_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, 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_inc(x_13); -x_16 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__6(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__5(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__7(x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_17); +x_18 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_evalTactic___spec__6(x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_17); lean_dec(x_13); lean_dec(x_10); lean_dec(x_8); return x_18; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -7260,7 +7086,7 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19, lean_object* x_20) { _start: { double x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; @@ -7273,7 +7099,7 @@ lean_ctor_set(x_22, 1, x_3); lean_ctor_set_float(x_22, sizeof(void*)*2, x_21); lean_ctor_set_float(x_22, sizeof(void*)*2 + 8, x_21); lean_ctor_set_uint8(x_22, sizeof(void*)*2 + 16, x_2); -x_23 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1; +x_23 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2___closed__1; x_24 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_23); if (x_24 == 0) { @@ -7283,7 +7109,7 @@ lean_object* x_25; lean_object* x_26; lean_dec(x_3); lean_dec(x_1); x_25 = lean_box(0); -x_26 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_4, x_5, x_11, x_6, x_22, x_25, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_26 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(x_4, x_5, x_11, x_6, x_22, x_25, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_6); return x_26; } @@ -7298,7 +7124,7 @@ lean_ctor_set_float(x_27, sizeof(void*)*2, x_9); lean_ctor_set_float(x_27, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_27, sizeof(void*)*2 + 16, x_2); x_28 = lean_box(0); -x_29 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_4, x_5, x_11, x_6, x_27, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_29 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(x_4, x_5, x_11, x_6, x_27, x_28, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_6); return x_29; } @@ -7314,13 +7140,13 @@ lean_ctor_set_float(x_30, sizeof(void*)*2, x_9); lean_ctor_set_float(x_30, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_30, sizeof(void*)*2 + 16, x_2); x_31 = lean_box(0); -x_32 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__2(x_4, x_5, x_11, x_6, x_30, x_31, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); +x_32 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__1(x_4, x_5, x_11, x_6, x_30, x_31, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_20); lean_dec(x_6); return x_32; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_evalTactic___spec__3___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -7328,16 +7154,16 @@ x_1 = lean_mk_string_from_bytes(" #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -85,7 +85,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalChoice_declRange___close static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalDone_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalChoiceAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalCase_x27___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*); -extern lean_object* l_Lean_diag; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRotateRight___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalChoice___closed__4; @@ -286,6 +285,7 @@ lean_object* l_Lean_addBuiltinDeclarationRanges(lean_object*, lean_object*, lean static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRange(lean_object*); static lean_object* l_Lean_Elab_Tactic_addCheckpoints___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_OpenDecl_elabOpenDecl___at_Lean_Elab_Tactic_evalOpen___spec__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalChoice(lean_object*); @@ -392,6 +392,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_addCheckpoints_ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalReplace___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalOpen___spec__21(lean_object*, 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_evalChoice_declRange___closed__6; +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); lean_object* l_Lean_Syntax_getNumArgs(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_renameInaccessibles___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*); @@ -403,6 +404,8 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRevert_declRange___close static lean_object* l_Lean_Elab_Tactic_evalFail___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalTacticSeqBracketed___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___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___lambda__1(lean_object*, uint8_t, 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 l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CommandContextInfo_saveNoFileMap___at_Lean_Elab_Tactic_renameInaccessibles___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_renameInaccessibles___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRefl(lean_object*); @@ -442,10 +445,10 @@ LEAN_EXPORT lean_object* l_Lean_pushScope___at_Lean_Elab_Tactic_evalOpen___spec_ lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRight___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_throwUnsupportedSyntax___at_Lean_Elab_Tactic_evalFirst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ensureNoOverload___at_Lean_Elab_Tactic_evalOpen___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_EXPORT lean_object* l_Lean_Elab_Tactic_evalAnyGoals___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntros___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_declRange___closed__5; @@ -802,6 +805,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange___close static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalFail___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRevert___closed__2; static lean_object* l___private_Lean_Elab_Open_0__Lean_Elab_OpenDecl_resolveNameUsingNamespacesCore___at_Lean_Elab_Tactic_evalOpen___spec__34___lambda__1___closed__2; +extern lean_object* l_Lean_diagnostics; static lean_object* l___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_getCaseGoals_notFound___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__42(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* l_Lean_LocalContext_getAt_x3f(lean_object*, lean_object*); @@ -1069,6 +1073,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_addCheckpoints___lambda__1___boxed(l LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntros___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___regBuiltin_Lean_Elab_Tactic_evalSkip(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalIntro___closed__27; +static lean_object* l_Lean_Elab_Tactic_elabSetOption___closed__3; 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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRange___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalTacticCDot___closed__2; @@ -1153,6 +1158,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRange___ lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_BuiltinTactic_0__Lean_Elab_Tactic_getCaseGoals_commaList___closed__5; lean_object* l_List_toString___at_Lean_ensureNoOverload___spec__2(lean_object*); +static lean_object* l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1; lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalContradiction_declRange___closed__7; static lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_Elab_Tactic_elabSetOption___spec__6___closed__1; @@ -1627,7 +1633,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1639,7 +1645,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(26u); +x_1 = lean_unsigned_to_nat(27u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1667,7 +1673,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1679,7 +1685,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalWithAnnotateState_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(23u); +x_1 = lean_unsigned_to_nat(24u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1825,7 +1831,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDone_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(28u); +x_1 = lean_unsigned_to_nat(29u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1837,7 +1843,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDone_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(29u); +x_1 = lean_unsigned_to_nat(30u); x_2 = lean_unsigned_to_nat(6u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1865,7 +1871,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDone_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(28u); +x_1 = lean_unsigned_to_nat(29u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1877,7 +1883,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDone_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(28u); +x_1 = lean_unsigned_to_nat(29u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2396,7 +2402,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(31u); +x_1 = lean_unsigned_to_nat(32u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2407,12 +2413,13 @@ return x_3; static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(37u); -x_2 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_2, 0, x_1); -lean_ctor_set(x_2, 1, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_unsigned_to_nat(38u); +x_2 = lean_unsigned_to_nat(37u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange___closed__3() { @@ -2435,7 +2442,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(31u); +x_1 = lean_unsigned_to_nat(32u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2447,7 +2454,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSeq1_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(31u); +x_1 = lean_unsigned_to_nat(32u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2576,7 +2583,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(39u); +x_1 = lean_unsigned_to_nat(40u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2588,7 +2595,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(40u); +x_1 = lean_unsigned_to_nat(41u); x_2 = lean_unsigned_to_nat(19u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2616,7 +2623,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(39u); +x_1 = lean_unsigned_to_nat(40u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -2628,7 +2635,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalParen_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(39u); +x_1 = lean_unsigned_to_nat(40u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3869,7 +3876,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3881,7 +3888,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(116u); +x_1 = lean_unsigned_to_nat(117u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3909,7 +3916,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -3921,7 +3928,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq1Indented _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(115u); +x_1 = lean_unsigned_to_nat(116u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4189,7 +4196,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeqBracketed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(118u); +x_1 = lean_unsigned_to_nat(119u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4201,7 +4208,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeqBracketed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(123u); +x_1 = lean_unsigned_to_nat(124u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4229,7 +4236,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeqBracketed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(118u); +x_1 = lean_unsigned_to_nat(119u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4241,7 +4248,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeqBracketed _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(118u); +x_1 = lean_unsigned_to_nat(119u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4431,7 +4438,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticCDot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(125u); +x_1 = lean_unsigned_to_nat(126u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4443,7 +4450,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticCDot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(134u); +x_1 = lean_unsigned_to_nat(135u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4471,7 +4478,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticCDot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(125u); +x_1 = lean_unsigned_to_nat(126u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4483,7 +4490,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticCDot_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(125u); +x_1 = lean_unsigned_to_nat(126u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4681,7 +4688,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4693,7 +4700,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(141u); +x_1 = lean_unsigned_to_nat(142u); x_2 = lean_unsigned_to_nat(21u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4721,7 +4728,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4733,7 +4740,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFocus_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(136u); +x_1 = lean_unsigned_to_nat(137u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4911,7 +4918,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateLeft_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(146u); +x_1 = lean_unsigned_to_nat(147u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4923,7 +4930,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateLeft_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(148u); +x_1 = lean_unsigned_to_nat(149u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4951,7 +4958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateLeft_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(146u); +x_1 = lean_unsigned_to_nat(147u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -4963,7 +4970,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateLeft_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(146u); +x_1 = lean_unsigned_to_nat(147u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5110,7 +5117,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(150u); +x_1 = lean_unsigned_to_nat(151u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5122,7 +5129,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(152u); +x_1 = lean_unsigned_to_nat(153u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5150,7 +5157,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(150u); +x_1 = lean_unsigned_to_nat(151u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5162,7 +5169,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRotateRight_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(150u); +x_1 = lean_unsigned_to_nat(151u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12460,7 +12467,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(154u); +x_1 = lean_unsigned_to_nat(155u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12472,7 +12479,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(162u); +x_1 = lean_unsigned_to_nat(163u); x_2 = lean_unsigned_to_nat(12u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12500,7 +12507,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(154u); +x_1 = lean_unsigned_to_nat(155u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -12512,7 +12519,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalOpen_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(154u); +x_1 = lean_unsigned_to_nat(155u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13229,7 +13236,7 @@ return x_73; } } } -static lean_object* _init_l_Lean_Elab_Tactic_elabSetOption___closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1() { _start: { lean_object* x_1; @@ -13237,11 +13244,449 @@ x_1 = l_Lean_maxRecDepth; return x_1; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___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, uint8_t x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_ctor_get(x_12, 2); +lean_dec(x_17); +x_18 = l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1; +x_19 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_18); +lean_ctor_set(x_12, 4, x_19); +lean_ctor_set(x_12, 2, x_1); +lean_ctor_set_uint8(x_12, sizeof(void*)*11 + 1, x_2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_9); +lean_inc(x_8); +x_20 = l_Lean_Elab_Tactic_evalTactic(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_12, x_13, x_14); +if (lean_obj_tag(x_20) == 0) +{ +if (x_10 == 0) +{ +uint8_t x_21; +lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) +{ +return x_20; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_20, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); +lean_dec(x_20); +x_27 = l_Lean_Meta_reportDiag(x_8, x_9, x_12, x_13, x_26); +if (lean_obj_tag(x_27) == 0) +{ +uint8_t x_28; +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) +{ +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set(x_27, 0, x_25); +return x_27; +} +else +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} +} +else +{ +uint8_t x_32; +lean_dec(x_25); +x_32 = !lean_is_exclusive(x_27); +if (x_32 == 0) +{ +return x_27; +} +else +{ +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_27, 0); +x_34 = lean_ctor_get(x_27, 1); +lean_inc(x_34); +lean_inc(x_33); +lean_dec(x_27); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; +} +} +} +} +else +{ +if (x_10 == 0) +{ +uint8_t x_36; +lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +x_36 = !lean_is_exclusive(x_20); +if (x_36 == 0) +{ +return x_20; +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_37 = lean_ctor_get(x_20, 0); +x_38 = lean_ctor_get(x_20, 1); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_20); +x_39 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +return x_39; +} +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_20, 0); +lean_inc(x_40); +x_41 = lean_ctor_get(x_20, 1); +lean_inc(x_41); +lean_dec(x_20); +x_42 = l_Lean_Meta_reportDiag(x_8, x_9, x_12, x_13, x_41); +if (lean_obj_tag(x_42) == 0) +{ +uint8_t x_43; +x_43 = !lean_is_exclusive(x_42); +if (x_43 == 0) +{ +lean_object* x_44; +x_44 = lean_ctor_get(x_42, 0); +lean_dec(x_44); +lean_ctor_set_tag(x_42, 1); +lean_ctor_set(x_42, 0, x_40); +return x_42; +} +else +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_42, 1); +lean_inc(x_45); +lean_dec(x_42); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_40); +lean_ctor_set(x_46, 1, x_45); +return x_46; +} +} +else +{ +uint8_t x_47; +lean_dec(x_40); +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_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_51 = lean_ctor_get(x_12, 0); +x_52 = lean_ctor_get(x_12, 1); +x_53 = lean_ctor_get(x_12, 3); +x_54 = lean_ctor_get(x_12, 5); +x_55 = lean_ctor_get(x_12, 6); +x_56 = lean_ctor_get(x_12, 7); +x_57 = lean_ctor_get(x_12, 8); +x_58 = lean_ctor_get(x_12, 9); +x_59 = lean_ctor_get(x_12, 10); +x_60 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); +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_inc(x_52); +lean_inc(x_51); +lean_dec(x_12); +x_61 = l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1; +x_62 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_61); +x_63 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_63, 0, x_51); +lean_ctor_set(x_63, 1, x_52); +lean_ctor_set(x_63, 2, x_1); +lean_ctor_set(x_63, 3, x_53); +lean_ctor_set(x_63, 4, x_62); +lean_ctor_set(x_63, 5, x_54); +lean_ctor_set(x_63, 6, x_55); +lean_ctor_set(x_63, 7, x_56); +lean_ctor_set(x_63, 8, x_57); +lean_ctor_set(x_63, 9, x_58); +lean_ctor_set(x_63, 10, x_59); +lean_ctor_set_uint8(x_63, sizeof(void*)*11, x_60); +lean_ctor_set_uint8(x_63, sizeof(void*)*11 + 1, x_2); +lean_inc(x_13); +lean_inc(x_63); +lean_inc(x_9); +lean_inc(x_8); +x_64 = l_Lean_Elab_Tactic_evalTactic(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_63, x_13, x_14); +if (lean_obj_tag(x_64) == 0) +{ +if (x_10 == 0) +{ +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_63); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_67 = x_64; +} else { + lean_dec_ref(x_64); + x_67 = lean_box(0); +} +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(0, 2, 0); +} else { + x_68 = x_67; +} +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +else +{ +lean_object* x_69; lean_object* x_70; lean_object* x_71; +x_69 = lean_ctor_get(x_64, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_64, 1); +lean_inc(x_70); +lean_dec(x_64); +x_71 = l_Lean_Meta_reportDiag(x_8, x_9, x_63, x_13, x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +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); +} +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(0, 2, 0); +} else { + x_74 = x_73; +} +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_72); +return x_74; +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_69); +x_75 = lean_ctor_get(x_71, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_71, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_77 = x_71; +} else { + lean_dec_ref(x_71); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; +} +} +} +else +{ +if (x_10 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +lean_dec(x_63); +lean_dec(x_13); +lean_dec(x_9); +lean_dec(x_8); +x_79 = lean_ctor_get(x_64, 0); +lean_inc(x_79); +x_80 = lean_ctor_get(x_64, 1); +lean_inc(x_80); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_81 = x_64; +} else { + lean_dec_ref(x_64); + x_81 = lean_box(0); +} +if (lean_is_scalar(x_81)) { + x_82 = lean_alloc_ctor(1, 2, 0); +} else { + x_82 = x_81; +} +lean_ctor_set(x_82, 0, x_79); +lean_ctor_set(x_82, 1, x_80); +return x_82; +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_64, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_64, 1); +lean_inc(x_84); +lean_dec(x_64); +x_85 = l_Lean_Meta_reportDiag(x_8, x_9, x_63, x_13, x_84); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +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); +} +if (lean_is_scalar(x_87)) { + x_88 = lean_alloc_ctor(1, 2, 0); +} else { + x_88 = x_87; + lean_ctor_set_tag(x_88, 1); +} +lean_ctor_set(x_88, 0, x_83); +lean_ctor_set(x_88, 1, x_86); +return x_88; +} +else +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +lean_dec(x_83); +x_89 = lean_ctor_get(x_85, 0); +lean_inc(x_89); +x_90 = lean_ctor_get(x_85, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_91 = x_85; +} else { + lean_dec_ref(x_85); + x_91 = lean_box(0); +} +if (lean_is_scalar(x_91)) { + x_92 = lean_alloc_ctor(1, 2, 0); +} else { + x_92 = x_91; +} +lean_ctor_set(x_92, 0, x_89); +lean_ctor_set(x_92, 1, x_90); +return x_92; +} +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabSetOption___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("diagnostics", 11); +return x_1; +} +} static lean_object* _init_l_Lean_Elab_Tactic_elabSetOption___closed__2() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Elab_Tactic_elabSetOption___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_elabSetOption___closed__3() { +_start: +{ lean_object* x_1; -x_1 = l_Lean_diag; +x_1 = l_Lean_diagnostics; return x_1; } } @@ -13261,10 +13706,11 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); +lean_inc(x_12); x_15 = l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1(x_12, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_15) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +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; uint8_t x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); @@ -13272,72 +13718,164 @@ lean_inc(x_17); lean_dec(x_15); x_18 = lean_unsigned_to_nat(5u); x_19 = l_Lean_Syntax_getArg(x_1, x_18); -x_20 = !lean_is_exclusive(x_8); -if (x_20 == 0) +lean_dec(x_1); +x_20 = l_Lean_Syntax_getId(x_12); +lean_dec(x_12); +x_21 = l_Lean_Elab_Tactic_elabSetOption___closed__2; +x_22 = lean_name_eq(x_20, x_21); +lean_dec(x_20); +x_23 = l_Lean_Elab_Tactic_elabSetOption___closed__3; +x_24 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_16, x_23); +x_25 = lean_st_ref_get(x_9, x_17); +if (x_22 == 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; lean_object* x_27; -x_21 = lean_ctor_get(x_8, 4); -lean_dec(x_21); -x_22 = lean_ctor_get(x_8, 2); -lean_dec(x_22); -x_23 = l_Lean_Elab_Tactic_elabSetOption___closed__1; -x_24 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_16, x_23); -x_25 = l_Lean_Elab_Tactic_elabSetOption___closed__2; -x_26 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_16, x_25); -lean_ctor_set(x_8, 4, x_24); -lean_ctor_set(x_8, 2, x_16); -lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_26); -x_27 = l_Lean_Elab_Tactic_evalTactic(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -return x_27; +lean_object* x_65; lean_object* x_66; uint8_t x_67; +x_65 = lean_ctor_get(x_25, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_25, 1); +lean_inc(x_66); +lean_dec(x_25); +x_67 = 0; +x_26 = x_67; +x_27 = x_65; +x_28 = x_66; +goto block_64; } 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; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; -x_28 = lean_ctor_get(x_8, 0); -x_29 = lean_ctor_get(x_8, 1); -x_30 = lean_ctor_get(x_8, 3); -x_31 = lean_ctor_get(x_8, 5); -x_32 = lean_ctor_get(x_8, 6); -x_33 = lean_ctor_get(x_8, 7); -x_34 = lean_ctor_get(x_8, 8); -x_35 = lean_ctor_get(x_8, 9); -x_36 = lean_ctor_get(x_8, 10); -x_37 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -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_object* x_68; lean_object* x_69; uint8_t x_70; +x_68 = lean_ctor_get(x_25, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_25, 1); +lean_inc(x_69); +lean_dec(x_25); +x_70 = 1; +x_26 = x_70; +x_27 = x_68; +x_28 = x_69; +goto block_64; +} +block_64: +{ +lean_object* x_29; uint8_t x_30; uint8_t x_31; +x_29 = lean_ctor_get(x_27, 0); lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_8); -x_38 = l_Lean_Elab_Tactic_elabSetOption___closed__1; -x_39 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_16, x_38); -x_40 = l_Lean_Elab_Tactic_elabSetOption___closed__2; -x_41 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_16, x_40); -x_42 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_42, 0, x_28); -lean_ctor_set(x_42, 1, x_29); -lean_ctor_set(x_42, 2, x_16); -lean_ctor_set(x_42, 3, x_30); -lean_ctor_set(x_42, 4, x_39); -lean_ctor_set(x_42, 5, x_31); -lean_ctor_set(x_42, 6, x_32); -lean_ctor_set(x_42, 7, x_33); -lean_ctor_set(x_42, 8, x_34); -lean_ctor_set(x_42, 9, x_35); -lean_ctor_set(x_42, 10, x_36); -lean_ctor_set_uint8(x_42, sizeof(void*)*11, x_37); -lean_ctor_set_uint8(x_42, sizeof(void*)*11 + 1, x_41); -x_43 = l_Lean_Elab_Tactic_evalTactic(x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_42, x_9, x_17); +lean_dec(x_27); +x_30 = l_Lean_Kernel_isDiagnosticsEnabled(x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +if (x_24 == 0) +{ +uint8_t x_60; +x_60 = 1; +x_31 = x_60; +goto block_59; +} +else +{ +uint8_t x_61; +x_61 = 0; +x_31 = x_61; +goto block_59; +} +} +else +{ +if (x_24 == 0) +{ +uint8_t x_62; +x_62 = 0; +x_31 = x_62; +goto block_59; +} +else +{ +uint8_t x_63; +x_63 = 1; +x_31 = x_63; +goto block_59; +} +} +block_59: +{ +if (x_31 == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_st_ref_take(x_9, x_28); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = !lean_is_exclusive(x_33); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_36 = lean_ctor_get(x_33, 0); +x_37 = lean_ctor_get(x_33, 4); +lean_dec(x_37); +x_38 = l_Lean_Kernel_enableDiag(x_36, x_24); +x_39 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__2___closed__4; +lean_ctor_set(x_33, 4, x_39); +lean_ctor_set(x_33, 0, x_38); +x_40 = lean_st_ref_set(x_9, x_33, x_34); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_box(0); +x_43 = l_Lean_Elab_Tactic_elabSetOption___lambda__1(x_16, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_26, x_42, x_8, x_9, 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; +x_44 = lean_ctor_get(x_33, 0); +x_45 = lean_ctor_get(x_33, 1); +x_46 = lean_ctor_get(x_33, 2); +x_47 = lean_ctor_get(x_33, 3); +x_48 = lean_ctor_get(x_33, 5); +x_49 = lean_ctor_get(x_33, 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_33); +x_50 = l_Lean_Kernel_enableDiag(x_44, x_24); +x_51 = l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__2___closed__4; +x_52 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_45); +lean_ctor_set(x_52, 2, x_46); +lean_ctor_set(x_52, 3, x_47); +lean_ctor_set(x_52, 4, x_51); +lean_ctor_set(x_52, 5, x_48); +lean_ctor_set(x_52, 6, x_49); +x_53 = lean_st_ref_set(x_9, x_52, x_34); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = lean_box(0); +x_56 = l_Lean_Elab_Tactic_elabSetOption___lambda__1(x_16, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_26, x_55, x_8, x_9, x_54); +return x_56; +} } else { -uint8_t x_44; +lean_object* x_57; lean_object* x_58; +x_57 = lean_box(0); +x_58 = l_Lean_Elab_Tactic_elabSetOption___lambda__1(x_16, x_24, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_26, x_57, x_8, x_9, x_28); +return x_58; +} +} +} +} +else +{ +uint8_t x_71; +lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -13346,23 +13884,24 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_44 = !lean_is_exclusive(x_15); -if (x_44 == 0) +lean_dec(x_1); +x_71 = !lean_is_exclusive(x_15); +if (x_71 == 0) { return x_15; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_15, 0); -x_46 = lean_ctor_get(x_15, 1); -lean_inc(x_46); -lean_inc(x_45); +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_15, 0); +x_73 = lean_ctor_get(x_15, 1); +lean_inc(x_73); +lean_inc(x_72); lean_dec(x_15); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -return x_47; +x_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; } } } @@ -13462,13 +14001,17 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_elabSetOption___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_elabSetOption___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_11; -x_11 = l_Lean_Elab_Tactic_elabSetOption(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_1); -return x_11; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = lean_unbox(x_2); +lean_dec(x_2); +x_16 = lean_unbox(x_10); +lean_dec(x_10); +x_17 = l_Lean_Elab_Tactic_elabSetOption___lambda__1(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_16, x_11, x_12, x_13, x_14); +lean_dec(x_11); +return x_17; } } static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__1() { @@ -13515,7 +14058,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabSetOption___boxed), 10, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabSetOption), 10, 0); return x_1; } } @@ -13535,7 +14078,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(164u); +x_1 = lean_unsigned_to_nat(165u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13547,8 +14090,8 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(167u); -x_2 = lean_unsigned_to_nat(21u); +x_1 = lean_unsigned_to_nat(172u); +x_2 = lean_unsigned_to_nat(18u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -13562,7 +14105,7 @@ lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_obj x_1 = l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRange___closed__1; x_2 = lean_unsigned_to_nat(43u); x_3 = l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRange___closed__2; -x_4 = lean_unsigned_to_nat(21u); +x_4 = lean_unsigned_to_nat(18u); x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -13575,7 +14118,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(164u); +x_1 = lean_unsigned_to_nat(165u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13587,7 +14130,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(164u); +x_1 = lean_unsigned_to_nat(165u); x_2 = lean_unsigned_to_nat(60u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14533,7 +15076,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAllGoals_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14545,7 +15088,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAllGoals_declRange _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(189u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14573,7 +15116,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAllGoals_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14585,7 +15128,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAllGoals_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(169u); +x_1 = lean_unsigned_to_nat(174u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15427,7 +15970,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange _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(191u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15439,7 +15982,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange _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(206u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15467,7 +16010,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange _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(191u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15479,7 +16022,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAnyGoals_declRange _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(191u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15588,7 +16131,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq_declRang _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(208u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15600,7 +16143,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq_declRang _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(209u); x_2 = lean_unsigned_to_nat(19u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15628,7 +16171,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq_declRang _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(208u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15640,7 +16183,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTacticSeq_declRang _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(208u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16355,7 +16898,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalChoice_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(220u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16367,7 +16910,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalChoice_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(221u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16395,7 +16938,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalChoice_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(220u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16407,7 +16950,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalChoice_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(220u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16553,7 +17096,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSkip_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(223u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16565,7 +17108,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSkip_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(223u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16593,7 +17136,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSkip_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(223u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16605,7 +17148,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSkip_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(223u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16748,7 +17291,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalUnknown_declRange_ _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(225u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16760,7 +17303,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalUnknown_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(226u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16788,7 +17331,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalUnknown_declRange_ _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(225u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16800,7 +17343,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalUnknown_declRange_ _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(225u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17159,7 +17702,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_decl _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(228u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17171,7 +17714,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_decl _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(232u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17199,7 +17742,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_decl _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(228u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17211,7 +17754,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFailIfSuccess_decl _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(228u); x_2 = lean_unsigned_to_nat(53u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17483,7 +18026,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceState_declRan _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(234u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17495,7 +18038,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceState_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(236u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17523,7 +18066,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceState_declRan _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(234u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17535,7 +18078,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceState_declRan _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(234u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17802,7 +18345,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceMessage_declR _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(238u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17814,7 +18357,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceMessage_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(236u); +x_1 = lean_unsigned_to_nat(241u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17842,7 +18385,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceMessage_declR _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(238u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -17854,7 +18397,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalTraceMessage_declR _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(238u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18308,7 +18851,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAssumption_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(243u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18320,7 +18863,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAssumption_declRan _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(247u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18348,7 +18891,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAssumption_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(243u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18360,7 +18903,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalAssumption_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(243u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18666,7 +19209,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalContradiction_decl _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(249u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18678,7 +19221,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalContradiction_decl _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(250u); x_2 = lean_unsigned_to_nat(63u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18706,7 +19249,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalContradiction_decl _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(249u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18718,7 +19261,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalContradiction_decl _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(249u); x_2 = lean_unsigned_to_nat(72u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19008,7 +19551,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRefl_declRange___c _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(252u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19020,7 +19563,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRefl_declRange___c _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(253u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19048,7 +19591,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRefl_declRange___c _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(252u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -19060,7 +19603,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRefl_declRange___c _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(252u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21013,7 +21556,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro_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(255u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21025,7 +21568,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(276u); +x_1 = lean_unsigned_to_nat(281u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21053,7 +21596,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro_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(255u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21065,7 +21608,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntro_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(255u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21324,7 +21867,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRan _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(283u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21336,7 +21879,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRan _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(286u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21364,7 +21907,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRan _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(283u); x_2 = lean_unsigned_to_nat(52u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -21376,7 +21919,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntroMatch_declRan _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(283u); x_2 = lean_unsigned_to_nat(66u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22241,7 +22784,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntros_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(288u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22253,7 +22796,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntros_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(295u); +x_1 = lean_unsigned_to_nat(300u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22281,7 +22824,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntros_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(288u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22293,7 +22836,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalIntros_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(288u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22588,7 +23131,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRevert_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(302u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22600,7 +23143,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRevert_declRange__ _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(307u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22628,7 +23171,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRevert_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(302u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -22640,7 +23183,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRevert_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(302u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23140,7 +23683,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClear_declRange___ _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(309u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23152,7 +23695,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClear_declRange___ _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(318u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23180,7 +23723,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClear_declRange___ _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(309u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23192,7 +23735,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalClear_declRange___ _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(309u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23670,7 +24213,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(322u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(43u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23682,7 +24225,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(325u); +x_1 = lean_unsigned_to_nat(330u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23710,7 +24253,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(322u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23722,7 +24265,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(322u); +x_1 = lean_unsigned_to_nat(327u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24017,7 +24560,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRang _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(332u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24029,7 +24572,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(328u); +x_1 = lean_unsigned_to_nat(333u); x_2 = lean_unsigned_to_nat(58u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24057,7 +24600,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRang _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(332u); x_2 = lean_unsigned_to_nat(51u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24069,7 +24612,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstVars_declRang _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(332u); x_2 = lean_unsigned_to_nat(64u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24334,7 +24877,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstEqs_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24346,7 +24889,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstEqs_declRange _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(336u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24374,7 +24917,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstEqs_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -24386,7 +24929,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSubstEqs_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(330u); +x_1 = lean_unsigned_to_nat(335u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32757,7 +33300,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(425u); +x_1 = lean_unsigned_to_nat(430u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32769,7 +33312,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(435u); +x_1 = lean_unsigned_to_nat(440u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32797,7 +33340,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(425u); +x_1 = lean_unsigned_to_nat(430u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -32809,7 +33352,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(425u); +x_1 = lean_unsigned_to_nat(430u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34303,7 +34846,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_x27_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(437u); +x_1 = lean_unsigned_to_nat(442u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34315,7 +34858,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_x27_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(452u); +x_1 = lean_unsigned_to_nat(457u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34343,7 +34886,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_x27_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(437u); +x_1 = lean_unsigned_to_nat(442u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34355,7 +34898,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalCase_x27_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(437u); +x_1 = lean_unsigned_to_nat(442u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34601,7 +35144,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessible _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(454u); +x_1 = lean_unsigned_to_nat(459u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34613,7 +35156,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessible _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(456u); +x_1 = lean_unsigned_to_nat(461u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34641,7 +35184,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessible _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(454u); +x_1 = lean_unsigned_to_nat(459u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -34653,7 +35196,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRenameInaccessible _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(454u); +x_1 = lean_unsigned_to_nat(459u); x_2 = lean_unsigned_to_nat(55u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35132,7 +35675,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFirst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(458u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35144,7 +35687,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFirst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(467u); +x_1 = lean_unsigned_to_nat(472u); x_2 = lean_unsigned_to_nat(48u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35172,7 +35715,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFirst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(458u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35184,7 +35727,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFirst_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(458u); +x_1 = lean_unsigned_to_nat(463u); x_2 = lean_unsigned_to_nat(47u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35487,7 +36030,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(469u); +x_1 = lean_unsigned_to_nat(474u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35499,7 +36042,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(475u); +x_1 = lean_unsigned_to_nat(480u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35527,7 +36070,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(469u); +x_1 = lean_unsigned_to_nat(474u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35539,7 +36082,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalFail_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(469u); +x_1 = lean_unsigned_to_nat(474u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35729,7 +36272,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(482u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35741,7 +36284,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(480u); +x_1 = lean_unsigned_to_nat(485u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35769,7 +36312,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(482u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35781,7 +36324,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalDbgTrace_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(477u); +x_1 = lean_unsigned_to_nat(482u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35952,7 +36495,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(482u); +x_1 = lean_unsigned_to_nat(487u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35964,7 +36507,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(485u); +x_1 = lean_unsigned_to_nat(490u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35992,7 +36535,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(482u); +x_1 = lean_unsigned_to_nat(487u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36004,7 +36547,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalSleep_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(482u); +x_1 = lean_unsigned_to_nat(487u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36318,7 +36861,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalLeft_declRange___c _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(492u); x_2 = lean_unsigned_to_nat(23u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36330,7 +36873,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalLeft_declRange___c _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(493u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36358,7 +36901,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalLeft_declRange___c _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(492u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36370,7 +36913,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalLeft_declRange___c _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(492u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36674,7 +37217,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRight_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(490u); +x_1 = lean_unsigned_to_nat(495u); x_2 = lean_unsigned_to_nat(24u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36686,7 +37229,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRight_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(491u); +x_1 = lean_unsigned_to_nat(496u); x_2 = lean_unsigned_to_nat(62u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36714,7 +37257,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRight_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(490u); +x_1 = lean_unsigned_to_nat(495u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -36726,7 +37269,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRight_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(490u); +x_1 = lean_unsigned_to_nat(495u); x_2 = lean_unsigned_to_nat(37u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37278,7 +37821,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalReplace_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(493u); +x_1 = lean_unsigned_to_nat(498u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37290,7 +37833,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalReplace_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(505u); +x_1 = lean_unsigned_to_nat(510u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37318,7 +37861,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalReplace_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(493u); +x_1 = lean_unsigned_to_nat(498u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37330,7 +37873,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalReplace_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(493u); +x_1 = lean_unsigned_to_nat(498u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37745,7 +38288,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(507u); +x_1 = lean_unsigned_to_nat(512u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37757,7 +38300,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(512u); +x_1 = lean_unsigned_to_nat(517u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37785,7 +38328,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(507u); +x_1 = lean_unsigned_to_nat(512u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37797,7 +38340,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_evalRunTac_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(507u); +x_1 = lean_unsigned_to_nat(512u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -37843,6 +38386,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } +lean_object* initialize_Lean_Meta_Diagnostics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Apply(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Assumption(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Contradiction(uint8_t builtin, lean_object*); @@ -37859,6 +38403,9 @@ LEAN_EXPORT lean_object* initialize_Lean_Elab_Tactic_BuiltinTactic(uint8_t built lean_object * res; if (_G_initialized) return lean_io_result_mk_ok(lean_box(0)); _G_initialized = true; +res = initialize_Lean_Meta_Diagnostics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Apply(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -38408,10 +38955,14 @@ l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed lean_mark_persistent(l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__5); l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__6 = _init_l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__6(); lean_mark_persistent(l_Lean_Elab_elabSetOption___at_Lean_Elab_Tactic_elabSetOption___spec__1___closed__6); +l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabSetOption___lambda__1___closed__1); l_Lean_Elab_Tactic_elabSetOption___closed__1 = _init_l_Lean_Elab_Tactic_elabSetOption___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSetOption___closed__1); l_Lean_Elab_Tactic_elabSetOption___closed__2 = _init_l_Lean_Elab_Tactic_elabSetOption___closed__2(); lean_mark_persistent(l_Lean_Elab_Tactic_elabSetOption___closed__2); +l_Lean_Elab_Tactic_elabSetOption___closed__3 = _init_l_Lean_Elab_Tactic_elabSetOption___closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_elabSetOption___closed__3); l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__1); l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__2 = _init_l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Pattern.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Pattern.c index 9a20630584ae..10f70ccdf85f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Pattern.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Pattern.c @@ -24,6 +24,7 @@ static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__1 uint32_t l_UInt32_ofNatTruncate(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalPattern(lean_object*); lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18; uint8_t l_Lean_Expr_isApp(lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__5; extern lean_object* l_Lean_Elab_Tactic_tacticElabAttribute; @@ -62,6 +63,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalPattern_declRange__ static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__6; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_Conv_evalPattern___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__13; +static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19; static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__9; lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___closed__3; @@ -196,6 +198,7 @@ lean_object* l_Lean_Elab_Tactic_replaceMainGoal(lean_object*, lean_object*, lean lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_Conv_evalPattern___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__10; static uint32_t _init_l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__1() { @@ -224,47 +227,51 @@ x_3 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_1, x_2); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) { -lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; uint32_t x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_5; lean_object* x_6; uint32_t x_7; lean_object* x_8; uint32_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; x_5 = lean_ctor_get(x_3, 0); x_6 = lean_box(0); x_7 = 0; x_8 = l_Lean_Meta_Simp_neutralConfig; x_9 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__1; x_10 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__2; -x_11 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_11, 0, x_8); -lean_ctor_set(x_11, 1, x_10); -lean_ctor_set(x_11, 2, x_5); -lean_ctor_set(x_11, 3, x_6); -lean_ctor_set_uint32(x_11, sizeof(void*)*4, x_9); -lean_ctor_set_uint32(x_11, sizeof(void*)*4 + 4, x_7); -lean_ctor_set(x_3, 0, x_11); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_10); +lean_ctor_set(x_12, 2, x_5); +lean_ctor_set(x_12, 3, x_6); +lean_ctor_set(x_12, 4, x_11); +lean_ctor_set_uint32(x_12, sizeof(void*)*5, x_9); +lean_ctor_set_uint32(x_12, sizeof(void*)*5 + 4, x_7); +lean_ctor_set(x_3, 0, x_12); return x_3; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint32_t x_15; lean_object* x_16; uint32_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_12 = lean_ctor_get(x_3, 0); -x_13 = lean_ctor_get(x_3, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint32_t x_16; lean_object* x_17; uint32_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_13 = lean_ctor_get(x_3, 0); +x_14 = lean_ctor_get(x_3, 1); +lean_inc(x_14); lean_inc(x_13); -lean_inc(x_12); lean_dec(x_3); -x_14 = lean_box(0); -x_15 = 0; -x_16 = l_Lean_Meta_Simp_neutralConfig; -x_17 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__1; -x_18 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__2; -x_19 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_19, 0, x_16); -lean_ctor_set(x_19, 1, x_18); -lean_ctor_set(x_19, 2, x_12); -lean_ctor_set(x_19, 3, x_14); -lean_ctor_set_uint32(x_19, sizeof(void*)*4, x_17); -lean_ctor_set_uint32(x_19, sizeof(void*)*4 + 4, x_15); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_13); -return x_20; +x_15 = lean_box(0); +x_16 = 0; +x_17 = l_Lean_Meta_Simp_neutralConfig; +x_18 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__1; +x_19 = l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_getContext___rarg___closed__2; +x_20 = lean_unsigned_to_nat(0u); +x_21 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_21, 0, x_17); +lean_ctor_set(x_21, 1, x_19); +lean_ctor_set(x_21, 2, x_13); +lean_ctor_set(x_21, 3, x_15); +lean_ctor_set(x_21, 4, x_20); +lean_ctor_set_uint32(x_21, sizeof(void*)*5, x_18); +lean_ctor_set_uint32(x_21, sizeof(void*)*5 + 4, x_16); +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_14); +return x_22; } } } @@ -3055,12 +3062,49 @@ return x_3; static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__6() { _start: { +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__2; +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_Elab_Tactic_Conv_evalPattern___lambda__9___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__5; +x_2 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__6; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__5; +x_2 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___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_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__9() { +_start: +{ lean_object* x_1; x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__5___boxed), 9, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__7() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__10() { _start: { lean_object* x_1; @@ -3068,16 +3112,16 @@ x_1 = lean_mk_string_from_bytes("'pattern' conv tactic failed, pattern was not f return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__8() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__7; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__10; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__9() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__12() { _start: { lean_object* x_1; @@ -3085,16 +3129,16 @@ x_1 = lean_mk_string_from_bytes("", 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__10() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__9; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__12; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__11() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__14() { _start: { lean_object* x_1; @@ -3102,16 +3146,16 @@ x_1 = lean_mk_string_from_bytes("'pattern' conv tactic failed, pattern was found return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__12() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__11; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__14; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__13() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16() { _start: { lean_object* x_1; @@ -3119,16 +3163,16 @@ x_1 = lean_mk_string_from_bytes(" times but ", 11); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__14() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__13; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18() { _start: { lean_object* x_1; @@ -3136,11 +3180,11 @@ x_1 = lean_mk_string_from_bytes(" expected", 9); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16() { +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15; +x_1 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -3148,7 +3192,7 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _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; uint32_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint32_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; uint8_t x_42; uint8_t x_43; uint8_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; uint8_t x_52; +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; uint32_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint32_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; uint8_t x_32; uint8_t x_33; uint8_t x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; uint8_t x_42; uint8_t x_43; uint8_t 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; uint8_t x_52; lean_object* x_53; uint8_t x_54; lean_inc(x_4); x_14 = lean_st_mk_ref(x_4, x_13); x_15 = lean_ctor_get(x_14, 0); @@ -3164,185 +3208,189 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_17, 1); lean_inc(x_20); lean_dec(x_17); -x_21 = lean_ctor_get_uint32(x_18, sizeof(void*)*4); +x_21 = lean_ctor_get_uint32(x_18, sizeof(void*)*5); x_22 = lean_ctor_get(x_18, 1); lean_inc(x_22); x_23 = lean_ctor_get(x_18, 2); lean_inc(x_23); x_24 = lean_ctor_get(x_18, 3); lean_inc(x_24); -x_25 = lean_ctor_get_uint32(x_18, sizeof(void*)*4 + 4); +x_25 = lean_ctor_get_uint32(x_18, sizeof(void*)*5 + 4); +x_26 = lean_ctor_get(x_18, 4); +lean_inc(x_26); if (lean_is_exclusive(x_18)) { lean_ctor_release(x_18, 0); lean_ctor_release(x_18, 1); lean_ctor_release(x_18, 2); lean_ctor_release(x_18, 3); - x_26 = x_18; + lean_ctor_release(x_18, 4); + x_27 = x_18; } else { lean_dec_ref(x_18); - x_26 = lean_box(0); + x_27 = lean_box(0); } -x_27 = lean_ctor_get(x_19, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_19, 1); +x_28 = lean_ctor_get(x_19, 0); lean_inc(x_28); -x_29 = lean_ctor_get_uint8(x_19, sizeof(void*)*2); -x_30 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 2); -x_31 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 3); -x_32 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 4); -x_33 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 5); -x_34 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 6); -x_35 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 7); -x_36 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 8); -x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 9); -x_38 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 10); -x_39 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 11); -x_40 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 12); -x_41 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 13); -x_42 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 14); -x_43 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 15); -x_44 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 16); +x_29 = lean_ctor_get(x_19, 1); +lean_inc(x_29); +x_30 = lean_ctor_get_uint8(x_19, sizeof(void*)*2); +x_31 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 2); +x_32 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 3); +x_33 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 4); +x_34 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 5); +x_35 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 6); +x_36 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 7); +x_37 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 8); +x_38 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 9); +x_39 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 10); +x_40 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 11); +x_41 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 12); +x_42 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 13); +x_43 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 14); +x_44 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 15); +x_45 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 16); if (lean_is_exclusive(x_19)) { lean_ctor_release(x_19, 0); lean_ctor_release(x_19, 1); - x_45 = x_19; + x_46 = x_19; } else { lean_dec_ref(x_19); - x_45 = lean_box(0); + x_46 = lean_box(0); } lean_inc(x_15); lean_inc(x_1); -x_46 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_pre___boxed), 11, 2); -lean_closure_set(x_46, 0, x_1); -lean_closure_set(x_46, 1, x_15); -lean_inc(x_2); -x_47 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__3___boxed), 10, 1); -lean_closure_set(x_47, 0, x_2); +x_47 = lean_alloc_closure((void*)(l___private_Lean_Elab_Tactic_Conv_Pattern_0__Lean_Elab_Tactic_Conv_pre___boxed), 11, 2); +lean_closure_set(x_47, 0, x_1); +lean_closure_set(x_47, 1, x_15); lean_inc(x_2); -x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__4___boxed), 10, 1); +x_48 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__3___boxed), 10, 1); lean_closure_set(x_48, 0, x_2); -x_49 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__6___boxed), 10, 1); +lean_inc(x_2); +x_49 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__4___boxed), 10, 1); lean_closure_set(x_49, 0, x_2); -x_50 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__6; -x_51 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_51, 0, x_46); -lean_ctor_set(x_51, 1, x_47); -lean_ctor_set(x_51, 2, x_48); -lean_ctor_set(x_51, 3, x_50); -lean_ctor_set(x_51, 4, x_49); +x_50 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__6___boxed), 10, 1); +lean_closure_set(x_50, 0, x_2); +x_51 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__9; +x_52 = 1; +x_53 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_53, 0, x_47); +lean_ctor_set(x_53, 1, x_48); +lean_ctor_set(x_53, 2, x_49); +lean_ctor_set(x_53, 3, x_51); +lean_ctor_set(x_53, 4, x_50); +lean_ctor_set_uint8(x_53, sizeof(void*)*5, x_52); if (lean_obj_tag(x_4) == 0) { -uint8_t x_127; lean_dec(x_4); -x_127 = 1; -x_52 = x_127; -goto block_126; +x_54 = x_52; +goto block_128; } else { -uint8_t x_128; +uint8_t x_129; lean_dec(x_4); -x_128 = 0; -x_52 = x_128; -goto block_126; +x_129 = 0; +x_54 = x_129; +goto block_128; } -block_126: +block_128: { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -if (lean_is_scalar(x_45)) { - x_53 = lean_alloc_ctor(0, 2, 17); +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +if (lean_is_scalar(x_46)) { + x_55 = lean_alloc_ctor(0, 2, 17); } else { - x_53 = x_45; -} -lean_ctor_set(x_53, 0, x_27); -lean_ctor_set(x_53, 1, x_28); -lean_ctor_set_uint8(x_53, sizeof(void*)*2, x_29); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 1, x_52); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 2, x_30); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 3, x_31); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 4, x_32); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 5, x_33); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 6, x_34); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 7, x_35); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 8, x_36); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 9, x_37); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 10, x_38); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 11, x_39); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 12, x_40); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 13, x_41); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 14, x_42); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 15, x_43); -lean_ctor_set_uint8(x_53, sizeof(void*)*2 + 16, x_44); -if (lean_is_scalar(x_26)) { - x_54 = lean_alloc_ctor(0, 4, 8); + x_55 = x_46; +} +lean_ctor_set(x_55, 0, x_28); +lean_ctor_set(x_55, 1, x_29); +lean_ctor_set_uint8(x_55, sizeof(void*)*2, x_30); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 1, x_54); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 2, x_31); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 3, x_32); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 4, x_33); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 5, x_34); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 6, x_35); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 7, x_36); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 8, x_37); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 9, x_38); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 10, x_39); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 11, x_40); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 12, x_41); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 13, x_42); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 14, x_43); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 15, x_44); +lean_ctor_set_uint8(x_55, sizeof(void*)*2 + 16, x_45); +if (lean_is_scalar(x_27)) { + x_56 = lean_alloc_ctor(0, 5, 8); } else { - x_54 = x_26; -} -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_22); -lean_ctor_set(x_54, 2, x_23); -lean_ctor_set(x_54, 3, x_24); -lean_ctor_set_uint32(x_54, sizeof(void*)*4, x_21); -lean_ctor_set_uint32(x_54, sizeof(void*)*4 + 4, x_25); -x_55 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__5; + x_56 = x_27; +} +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_22); +lean_ctor_set(x_56, 2, x_23); +lean_ctor_set(x_56, 3, x_24); +lean_ctor_set(x_56, 4, x_26); +lean_ctor_set_uint32(x_56, sizeof(void*)*5, x_21); +lean_ctor_set_uint32(x_56, sizeof(void*)*5 + 4, x_25); +x_57 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__8; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_56 = l_Lean_Meta_Simp_main(x_3, x_54, x_55, x_51, x_9, x_10, x_11, x_12, x_20); -if (lean_obj_tag(x_56) == 0) +x_58 = l_Lean_Meta_Simp_main(x_3, x_56, x_57, x_53, x_9, x_10, x_11, x_12, x_20); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); -lean_inc(x_58); -lean_dec(x_56); -x_59 = lean_ctor_get(x_57, 0); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_st_ref_get(x_15, x_58); -lean_dec(x_15); -x_61 = lean_ctor_get(x_60, 0); +x_60 = lean_ctor_get(x_58, 1); +lean_inc(x_60); +lean_dec(x_58); +x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); -if (lean_obj_tag(x_61) == 0) -{ -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_60, 1); -lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_ctor_get(x_61, 0); +lean_dec(x_59); +x_62 = lean_st_ref_get(x_15, x_60); +lean_dec(x_15); +x_63 = lean_ctor_get(x_62, 0); lean_inc(x_63); -lean_dec(x_61); -x_64 = lean_array_get_size(x_63); -x_65 = lean_unsigned_to_nat(0u); -x_66 = lean_nat_dec_eq(x_64, x_65); -lean_dec(x_64); -if (x_66 == 0) +if (lean_obj_tag(x_63) == 0) { -lean_object* x_67; +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_62, 1); +lean_inc(x_64); +lean_dec(x_62); +x_65 = lean_ctor_get(x_63, 0); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_array_get_size(x_65); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_eq(x_66, x_67); +lean_dec(x_66); +if (x_68 == 0) +{ +lean_object* x_69; lean_dec(x_1); -x_67 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__7(x_59, x_63, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); -return x_67; +x_69 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__7(x_61, x_65, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_64); +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; uint8_t x_75; -lean_dec(x_63); -lean_dec(x_59); -x_68 = lean_ctor_get(x_1, 2); -lean_inc(x_68); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +lean_dec(x_65); +lean_dec(x_61); +x_70 = lean_ctor_get(x_1, 2); +lean_inc(x_70); lean_dec(x_1); -x_69 = l_Lean_indentExpr(x_68); -x_70 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__8; -x_71 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_69); -x_72 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__10; +x_71 = l_Lean_indentExpr(x_70); +x_72 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__11; x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_71); -lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_throwError___at_Lean_Elab_Tactic_Conv_evalPattern___spec__2(x_73, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_62); +lean_ctor_set(x_73, 0, x_72); +lean_ctor_set(x_73, 1, x_71); +x_74 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__13; +x_75 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_75, 0, x_73); +lean_ctor_set(x_75, 1, x_74); +x_76 = l_Lean_throwError___at_Lean_Elab_Tactic_Conv_evalPattern___spec__2(x_75, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_64); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3351,96 +3399,96 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_75 = !lean_is_exclusive(x_74); -if (x_75 == 0) +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) { -return x_74; +return x_76; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_74, 0); -x_77 = lean_ctor_get(x_74, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_74); -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_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_76, 0); +x_79 = lean_ctor_get(x_76, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_76); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; uint8_t x_85; -x_79 = lean_ctor_get(x_60, 1); -lean_inc(x_79); -lean_dec(x_60); -x_80 = lean_ctor_get(x_61, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_61, 1); +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; uint8_t x_87; +x_81 = lean_ctor_get(x_62, 1); lean_inc(x_81); -x_82 = lean_ctor_get(x_61, 2); +lean_dec(x_62); +x_82 = lean_ctor_get(x_63, 0); lean_inc(x_82); -lean_dec(x_61); -x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__7), 11, 1); -lean_closure_set(x_83, 0, x_59); -x_84 = lean_unsigned_to_nat(0u); -x_85 = lean_nat_dec_eq(x_81, x_84); -if (x_85 == 0) +x_83 = lean_ctor_get(x_63, 1); +lean_inc(x_83); +x_84 = lean_ctor_get(x_63, 2); +lean_inc(x_84); +lean_dec(x_63); +x_85 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__7), 11, 1); +lean_closure_set(x_85, 0, x_61); +x_86 = lean_unsigned_to_nat(0u); +x_87 = lean_nat_dec_eq(x_83, x_86); +if (x_87 == 0) { -lean_object* x_86; +lean_object* x_88; lean_dec(x_1); -x_86 = l_List_getLast_x3f___rarg(x_82); -if (lean_obj_tag(x_86) == 0) +x_88 = l_List_getLast_x3f___rarg(x_84); +if (lean_obj_tag(x_88) == 0) { -lean_object* x_87; lean_object* x_88; -lean_dec(x_81); -x_87 = lean_box(0); -x_88 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__8(x_80, x_83, x_87, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_79); -return x_88; +lean_object* x_89; lean_object* x_90; +lean_dec(x_83); +x_89 = lean_box(0); +x_90 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__8(x_82, x_85, x_89, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); +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; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; -lean_dec(x_83); -lean_dec(x_80); -x_89 = lean_ctor_get(x_86, 0); -lean_inc(x_89); -lean_dec(x_86); -x_90 = lean_ctor_get(x_89, 0); -lean_inc(x_90); -lean_dec(x_89); -x_91 = l___private_Init_Data_Repr_0__Nat_reprFast(x_81); -x_92 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_92, 0, x_91); -x_93 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_93, 0, x_92); -x_94 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__12; -x_95 = lean_alloc_ctor(7, 2, 0); +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; +lean_dec(x_85); +lean_dec(x_82); +x_91 = lean_ctor_get(x_88, 0); +lean_inc(x_91); +lean_dec(x_88); +x_92 = lean_ctor_get(x_91, 0); +lean_inc(x_92); +lean_dec(x_91); +x_93 = l___private_Init_Data_Repr_0__Nat_reprFast(x_83); +x_94 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_95 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_93); -x_96 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__14; +x_96 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15; x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -x_98 = lean_unsigned_to_nat(1u); -x_99 = lean_nat_add(x_90, x_98); -lean_dec(x_90); -x_100 = l___private_Init_Data_Repr_0__Nat_reprFast(x_99); -x_101 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_101, 0, x_100); -x_102 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_102, 0, x_101); -x_103 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_103, 0, x_97); -lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16; +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_95); +x_98 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17; +x_99 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_99, 0, x_97); +lean_ctor_set(x_99, 1, x_98); +x_100 = lean_unsigned_to_nat(1u); +x_101 = lean_nat_add(x_92, x_100); +lean_dec(x_92); +x_102 = l___private_Init_Data_Repr_0__Nat_reprFast(x_101); +x_103 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_103, 0, x_102); +x_104 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_104, 0, x_103); x_105 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_105, 0, x_103); +lean_ctor_set(x_105, 0, x_99); lean_ctor_set(x_105, 1, x_104); -x_106 = l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic___spec__2(x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_79); +x_106 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19; +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 = l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic___spec__2(x_107, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3449,46 +3497,46 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_107 = !lean_is_exclusive(x_106); -if (x_107 == 0) +x_109 = !lean_is_exclusive(x_108); +if (x_109 == 0) { -return x_106; +return x_108; } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_108 = lean_ctor_get(x_106, 0); -x_109 = lean_ctor_get(x_106, 1); -lean_inc(x_109); -lean_inc(x_108); -lean_dec(x_106); -x_110 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_110, 0, x_108); -lean_ctor_set(x_110, 1, x_109); -return x_110; +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_108, 0); +x_111 = lean_ctor_get(x_108, 1); +lean_inc(x_111); +lean_inc(x_110); +lean_dec(x_108); +x_112 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_112, 0, x_110); +lean_ctor_set(x_112, 1, x_111); +return x_112; } } } else { -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; uint8_t x_118; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; uint8_t x_120; +lean_dec(x_85); +lean_dec(x_84); lean_dec(x_83); lean_dec(x_82); -lean_dec(x_81); -lean_dec(x_80); -x_111 = lean_ctor_get(x_1, 2); -lean_inc(x_111); +x_113 = lean_ctor_get(x_1, 2); +lean_inc(x_113); lean_dec(x_1); -x_112 = l_Lean_indentExpr(x_111); -x_113 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__8; -x_114 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_112); -x_115 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__10; +x_114 = l_Lean_indentExpr(x_113); +x_115 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__11; x_116 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_116, 0, x_114); -lean_ctor_set(x_116, 1, x_115); -x_117 = l_Lean_throwError___at_Lean_Elab_Tactic_Conv_evalPattern___spec__2(x_116, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_79); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +x_117 = l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__13; +x_118 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_118, 0, x_116); +lean_ctor_set(x_118, 1, x_117); +x_119 = l_Lean_throwError___at_Lean_Elab_Tactic_Conv_evalPattern___spec__2(x_118, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_81); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -3497,30 +3545,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_118 = !lean_is_exclusive(x_117); -if (x_118 == 0) +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) { -return x_117; +return x_119; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_117, 0); -x_120 = lean_ctor_get(x_117, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_117); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_119); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; } } } } else { -uint8_t x_122; +uint8_t x_124; lean_dec(x_15); lean_dec(x_12); lean_dec(x_11); @@ -3531,23 +3579,23 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_122 = !lean_is_exclusive(x_56); -if (x_122 == 0) +x_124 = !lean_is_exclusive(x_58); +if (x_124 == 0) { -return x_56; +return x_58; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_56, 0); -x_124 = lean_ctor_get(x_56, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_56); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_58, 0); +x_126 = lean_ctor_get(x_58, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_58); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } } } @@ -4621,6 +4669,12 @@ l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__15); l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__16); +l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__17); +l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__18); +l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__9___closed__19); l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__1); l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__2 = _init_l_Lean_Elab_Tactic_Conv_evalPattern___lambda__11___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c index a09388df061a..84b9ff7414b0 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Simp.c @@ -19,6 +19,7 @@ lean_object* l_Lean_Meta_Simp_Result_getProof(lean_object*, lean_object*, lean_o lean_object* l_Lean_Meta_getSimpTheorems___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; lean_object* l_Lean_Meta_dsimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange___closed__2; @@ -64,6 +65,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2(lean_objec static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalDSimp_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp___closed__10; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimpMatch_declRange(lean_object*); static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__2; @@ -98,6 +100,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1(lean_obje static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp___closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___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_Elab_Tactic_Conv_evalDSimp___closed__3; +static lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimpMatch___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_Conv_applySimpResult(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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: @@ -218,11 +221,48 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__2; +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_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5; +x_2 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5; +x_2 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__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; +} +} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; -x_14 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5; +x_14 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; x_15 = l_Lean_Meta_simp(x_1, x_2, x_3, x_4, x_14, x_9, x_10, x_11, x_12, x_13); return x_15; } @@ -993,7 +1033,7 @@ x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); x_21 = l_Lean_Elab_Tactic_Conv_evalDSimp___lambda__1___closed__1; -x_22 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5; +x_22 = l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -1308,6 +1348,12 @@ l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__4 = _init_l_Lean_Elab_Tac lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__4); l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__5); +l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__6); +l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__7); +l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8(); +lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__1___closed__8); l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1 = _init_l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Conv_evalSimp___lambda__2___closed__1); l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp___closed__1 = _init_l___regBuiltin_Lean_Elab_Tactic_Conv_evalSimp___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Ext.c b/stage0/stdlib/Lean/Elab/Tactic/Ext.c index fa43705b6e8e..1159293fd8d3 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Ext.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Ext.c @@ -472,6 +472,7 @@ lean_object* lean_string_length(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_614____lambda__9___boxed(lean_object*, lean_object*); lean_object* l_Lean_TSyntax_getNat(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +lean_object* l_Lean_throwError___at_Lean_MVarId_proofIrrelHeq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Ext_instReprExtTheorem; LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_Tactic_Ext_withExtN___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___hyg_614____lambda__7___closed__1; @@ -645,7 +646,6 @@ static lean_object* l_Lean_Elab_Tactic_Ext_initFn____x40_Lean_Elab_Tactic_Ext___ extern lean_object* l_Lean_Elab_unsupportedSyntaxExceptionId; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Ext_ExtTheorems_erase___spec__5(lean_object*, lean_object*, size_t, size_t, uint8_t); -lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_43____closed__22; lean_object* l_Lean_Meta_mkProjection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_Ext_0__Lean_Elab_Tactic_Ext_reprExtTheorem____x40_Lean_Elab_Tactic_Ext___hyg_43____boxed(lean_object*, lean_object*); @@ -7989,7 +7989,7 @@ x_9 = lean_ctor_get(x_5, 5); x_10 = l_Lean_replaceRef(x_1, x_9); lean_dec(x_9); lean_ctor_set(x_5, 5, x_10); -x_11 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(x_2, x_3, x_4, x_5, x_6, x_7); +x_11 = l_Lean_throwError___at_Lean_MVarId_proofIrrelHeq___spec__1(x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_5); return x_11; } @@ -8037,7 +8037,7 @@ lean_ctor_set(x_26, 9, x_21); lean_ctor_set(x_26, 10, x_22); lean_ctor_set_uint8(x_26, sizeof(void*)*11, x_23); lean_ctor_set_uint8(x_26, sizeof(void*)*11 + 1, x_24); -x_27 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(x_2, x_3, x_4, x_26, x_6, x_7); +x_27 = l_Lean_throwError___at_Lean_MVarId_proofIrrelHeq___spec__1(x_2, x_3, x_4, x_26, x_6, x_7); lean_dec(x_26); return x_27; } diff --git a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c index e839993bba57..ef27063f241d 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c +++ b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c @@ -32,7 +32,7 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_upwardAndElim___closed__2; static lean_object* l_Lean_Elab_Tactic_NormCast_normCastTarget___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsingDown___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkAppOptM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__5(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*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__10___closed__2; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__19; @@ -45,18 +45,19 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast_declRange___closed__4; static lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__9; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim_declRange___closed__5; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__11(lean_object*, lean_object*, lean_object*, uint32_t, lean_object*, lean_object*, uint32_t, 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*); uint32_t l_UInt32_ofNatTruncate(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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*); double lean_float_div(double, double); static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__10; lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__5; static lean_object* l_Lean_Elab_Tactic_NormCast_evalNormCast0___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___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_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__21; @@ -65,12 +66,10 @@ lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_obj LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_NormCast_evalNormCast0___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___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_Tactic_NormCast_derive___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___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*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__3; lean_object* l_Lean_Elab_Tactic_expandOptLocation(lean_object*); lean_object* l_Lean_Elab_Tactic_expandLocation(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__12; -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__12___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_upwardAndElim(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -82,9 +81,11 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1__ lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__8; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__8; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__10___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__11; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_prove___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13___boxed__const__1; @@ -96,7 +97,9 @@ extern lean_object* l_Lean_Expr_instBEq; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_numeralToCoe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast___closed__2; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___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_Lean_Elab_Tactic_NormCast_prove___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_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_getMainGoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_numeralToCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0(lean_object*); @@ -111,9 +114,10 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsingDown___lambda__ static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast_declRange___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__9; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___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_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__25; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___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_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_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*); @@ -134,11 +138,11 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___boxed( static lean_object* l_Lean_Elab_Tactic_NormCast_prove___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_prove___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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___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_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__2; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -148,12 +152,14 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__27; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__12___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRange___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13___closed__3; +static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0_declRange___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_numeralToCoe___lambda__1___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRange___closed__7; +static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__1; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__20; @@ -165,13 +171,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0_declR size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1___closed__6; 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_NormCast_derive___lambda__13___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast___closed__1; lean_object* l_Lean_Meta_findLocalDeclWithType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___closed__2; static uint32_t l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim___closed__3; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_prove___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0_declRange___closed__2; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -182,6 +189,7 @@ lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_p extern lean_object* l_Lean_Elab_Command_commandElabAttribute; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRange___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__13; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__7; uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__12; @@ -194,14 +202,15 @@ lean_object* l_Lean_Meta_Simp_Result_mkEqSymm(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_mkCoe(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabAuxDef___spec__1___rarg(lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabForall___spec__1___rarg(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__7; lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__3___boxed(lean_object**); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast___closed__5; +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(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*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_Conv_getLhs(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_NormCast_elabAddElim___closed__14; lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_Simp_discharge_x3f_x27___spec__7___rarg(lean_object*, lean_object*); @@ -209,9 +218,7 @@ lean_object* l_Lean_Meta_simp(lean_object*, lean_object*, lean_object*, lean_obj static lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___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_PersistentArray_append___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__32; lean_object* l_Lean_Meta_NormCast_addElim(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_prove___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast___closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_evalNormCast0___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -233,8 +240,7 @@ uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__11; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__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_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_NormCast_proveEqUsingDown___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___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_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1___closed__1; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__6; lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -258,11 +264,13 @@ lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__7; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__10___closed__4; extern lean_object* l_Lean_Meta_NormCast_pushCastExt; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__16; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13___closed__1; +extern lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__2; @@ -277,7 +285,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_isCoeOf_x3f___lambda__1(lea static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__26; static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_prove___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_Tactic_simpLocation(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_NormCast_elabModCast___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -303,8 +310,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__6(lean_obj LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___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_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_prove(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_NormCast_splittingProcedure___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_NormCast_normCastExt; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); @@ -321,17 +328,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRan static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalNormCast0_declRange(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_upwardAndElim___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___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_evalConvNormCast___boxed(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast(lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_evalConvNormCast___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___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast___closed__1; -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_prove___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_Tactic_NormCast_evalPushCast_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_evalPushCast___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___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange___closed__7; @@ -341,7 +345,7 @@ static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqU lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast___closed__4; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___lambda__4(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*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___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_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__2; static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__1___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__10; @@ -355,9 +359,11 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__28; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalPushCast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_normCastHyp(lean_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_NormCast_prove___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___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange___closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_Meta_coerce_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__5; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__13; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -369,7 +375,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_prove___lambda__2___boxed(l LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__5; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__5(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*); 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___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast_declRange___closed__4; lean_object* l_Lean_LocalDecl_type(lean_object*); @@ -388,7 +393,7 @@ lean_object* l_Lean_Meta_Simp_mkCongr(lean_object*, lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_normCastTarget___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast___closed__1; static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__10; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_NormCast_prove___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, 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_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__9; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkNumeral(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -403,9 +408,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__1___boxed( static lean_object* l_Lean_Elab_Tactic_NormCast_isNumeral_x3f___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast___closed__2; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__9; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsingDown___lambda__1___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure___lambda__3___closed__2; lean_object* l_Lean_exceptOptionEmoji___rarg(lean_object*); @@ -434,7 +439,7 @@ static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__13___closed__4; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6_(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim___closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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*); +static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__8; 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_object* l_Lean_Meta_Simp_mkDefaultMethods(lean_object*, lean_object*, lean_object*); @@ -445,12 +450,11 @@ lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_mkCoe___closed__1; static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__11; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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___regBuiltin_Lean_Elab_Tactic_NormCast_elabModCast(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_upwardAndElim___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_NormCast_isNumeral_x3f___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast_declRange___closed__5; @@ -460,19 +464,18 @@ lean_object* l_Lean_Meta_applySimpResultToTarget(lean_object*, lean_object*, lea lean_object* l_Lean_Elab_Tactic_Conv_applySimpResult(lean_object*, 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_NormCast_proveEqUsingDown___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_NormCast_evalNormCast0___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__4; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_prove___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__23; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim_declRange___closed__6; static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__22; +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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*); static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__12___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_evalConvNormCast___closed__3; static lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__8___closed__4; static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Tactic_NormCast_elabAddElim___closed__1; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_prove___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_normCastHyp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Tactic_replaceMainGoal(lean_object*, lean_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_NormCast_isNumeral_x3f(lean_object*); @@ -483,7 +486,7 @@ lean_object* lean_simp(lean_object*, lean_object*, lean_object*, lean_object*, l LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_splittingProcedure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__17; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_elabAddElim___closed__15; static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsingDown___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -492,6 +495,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_derive___lambda__9___boxed( static lean_object* l_Lean_Elab_Tactic_NormCast_initFn____x40_Lean_Elab_Tactic_NormCast___hyg_6____closed__18; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_splittingProcedure___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); double lean_float_sub(double, double); +static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__14; uint8_t l_Lean_Expr_hasExprMVar(lean_object*); static lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__6; static lean_object* l_Lean_Elab_Tactic_NormCast_evalNormCast0___closed__2; @@ -933,17 +937,44 @@ return x_3; static lean_object* _init_l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__12() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__6; +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_Elab_Tactic_NormCast_proveEqUsing___closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__11; +x_2 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__12; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__8; x_2 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__4; x_3 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__11; x_4 = lean_unsigned_to_nat(0u); -x_5 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_5, 0, x_1); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -lean_ctor_set(x_5, 3, x_4); -return x_5; +x_5 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__13; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_1); +lean_ctor_set(x_6, 1, x_2); +lean_ctor_set(x_6, 2, x_3); +lean_ctor_set(x_6, 3, x_4); +lean_ctor_set(x_6, 4, x_5); +return x_6; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_NormCast_proveEqUsing(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { @@ -958,7 +989,7 @@ x_10 = lean_ctor_get(x_4, 0); x_11 = !lean_is_exclusive(x_10); if (x_11 == 0) { -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; uint32_t x_22; lean_object* x_23; uint32_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; +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; uint32_t x_22; lean_object* x_23; uint32_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; x_12 = 2; lean_ctor_set_uint8(x_10, 9, x_12); x_13 = l_Lean_Meta_Simp_mkDefaultMethods(x_6, x_7, x_8); @@ -979,694 +1010,684 @@ x_21 = lean_box(0); x_22 = 0; x_23 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__1; x_24 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; -x_25 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_20); -lean_ctor_set(x_25, 2, x_17); -lean_ctor_set(x_25, 3, x_21); -lean_ctor_set_uint32(x_25, sizeof(void*)*4, x_24); -lean_ctor_set_uint32(x_25, sizeof(void*)*4 + 4, x_22); -x_26 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__12; -x_27 = lean_st_mk_ref(x_26, x_18); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +x_25 = lean_unsigned_to_nat(0u); +x_26 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_20); +lean_ctor_set(x_26, 2, x_17); +lean_ctor_set(x_26, 3, x_21); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set_uint32(x_26, sizeof(void*)*5, x_24); +lean_ctor_set_uint32(x_26, sizeof(void*)*5 + 4, x_22); +x_27 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__14; +x_28 = lean_st_mk_ref(x_27, x_18); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_28); -lean_inc(x_25); +lean_inc(x_29); +lean_inc(x_26); lean_inc(x_14); -x_30 = lean_simp(x_2, x_14, x_25, x_28, x_4, x_5, x_6, x_7, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = lean_simp(x_2, x_14, x_26, x_29, x_4, x_5, x_6, x_7, 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_28); -lean_inc(x_25); +lean_inc(x_29); +lean_inc(x_26); lean_inc(x_14); lean_inc(x_3); -x_33 = lean_simp(x_3, x_14, x_25, x_28, x_4, x_5, x_6, x_7, x_32); -if (lean_obj_tag(x_33) == 0) +x_34 = lean_simp(x_3, x_14, x_26, x_29, x_4, x_5, x_6, x_7, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_ctor_get(x_31, 0); +x_36 = lean_ctor_get(x_34, 1); lean_inc(x_36); -x_37 = lean_ctor_get(x_34, 0); +lean_dec(x_34); +x_37 = lean_ctor_get(x_32, 0); lean_inc(x_37); +x_38 = lean_ctor_get(x_35, 0); +lean_inc(x_38); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_38 = l_Lean_Meta_isExprDefEq(x_36, x_37, x_4, x_5, x_6, x_7, x_35); -if (lean_obj_tag(x_38) == 0) +x_39 = l_Lean_Meta_isExprDefEq(x_37, x_38, x_4, x_5, x_6, x_7, x_36); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_unbox(x_39); -lean_dec(x_39); -if (x_40 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; uint8_t x_43; -lean_dec(x_34); -lean_dec(x_31); -lean_dec(x_25); +lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_dec(x_35); +lean_dec(x_32); +lean_dec(x_26); lean_dec(x_14); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_41 = lean_ctor_get(x_38, 1); -lean_inc(x_41); -lean_dec(x_38); -x_42 = lean_st_ref_get(x_28, x_41); -lean_dec(x_28); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +x_42 = lean_ctor_get(x_39, 1); +lean_inc(x_42); +lean_dec(x_39); +x_43 = lean_st_ref_get(x_29, x_42); +lean_dec(x_29); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -lean_object* x_44; -x_44 = lean_ctor_get(x_42, 0); -lean_dec(x_44); -lean_ctor_set(x_42, 0, x_21); -return x_42; +lean_object* x_45; +x_45 = lean_ctor_get(x_43, 0); +lean_dec(x_45); +lean_ctor_set(x_43, 0, x_21); +return x_43; } else { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -lean_dec(x_42); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_21); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_object* x_46; lean_object* x_47; +x_46 = lean_ctor_get(x_43, 1); +lean_inc(x_46); +lean_dec(x_43); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_21); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_38, 1); -lean_inc(x_47); -lean_dec(x_38); -x_48 = lean_box(0); -x_49 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_34, x_31, x_48, x_14, x_25, x_28, x_4, x_5, x_6, x_7, x_47); -lean_dec(x_25); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = lean_box(0); +x_50 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_35, x_32, x_49, x_14, x_26, x_29, x_4, x_5, x_6, x_7, x_48); +lean_dec(x_26); lean_dec(x_14); -if (lean_obj_tag(x_49) == 0) +if (lean_obj_tag(x_50) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_49); -x_52 = lean_st_ref_get(x_28, x_51); -lean_dec(x_28); -x_53 = !lean_is_exclusive(x_52); -if (x_53 == 0) +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = lean_st_ref_get(x_29, x_52); +lean_dec(x_29); +x_54 = !lean_is_exclusive(x_53); +if (x_54 == 0) { -lean_object* x_54; -x_54 = lean_ctor_get(x_52, 0); -lean_dec(x_54); -lean_ctor_set(x_52, 0, x_50); -return x_52; +lean_object* x_55; +x_55 = lean_ctor_get(x_53, 0); +lean_dec(x_55); +lean_ctor_set(x_53, 0, x_51); +return x_53; } else { -lean_object* x_55; lean_object* x_56; -x_55 = lean_ctor_get(x_52, 1); -lean_inc(x_55); -lean_dec(x_52); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_50); -lean_ctor_set(x_56, 1, x_55); -return x_56; +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_53, 1); +lean_inc(x_56); +lean_dec(x_53); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_51); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } else { -uint8_t x_57; -lean_dec(x_28); -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_58; +lean_dec(x_29); +x_58 = !lean_is_exclusive(x_50); +if (x_58 == 0) { -return x_49; +return x_50; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_50, 0); +x_60 = lean_ctor_get(x_50, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_50); +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 { -uint8_t x_61; -lean_dec(x_34); -lean_dec(x_31); -lean_dec(x_28); -lean_dec(x_25); +uint8_t x_62; +lean_dec(x_35); +lean_dec(x_32); +lean_dec(x_29); +lean_dec(x_26); lean_dec(x_14); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_61 = !lean_is_exclusive(x_38); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_39); +if (x_62 == 0) { -return x_38; +return x_39; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_38, 0); -x_63 = lean_ctor_get(x_38, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_39, 0); +x_64 = lean_ctor_get(x_39, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_38); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_dec(x_39); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -uint8_t x_65; -lean_dec(x_31); -lean_dec(x_28); -lean_dec(x_25); +uint8_t x_66; +lean_dec(x_32); +lean_dec(x_29); +lean_dec(x_26); lean_dec(x_14); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_65 = !lean_is_exclusive(x_33); -if (x_65 == 0) +x_66 = !lean_is_exclusive(x_34); +if (x_66 == 0) { -return x_33; +return x_34; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_33, 0); -x_67 = lean_ctor_get(x_33, 1); +lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_34, 0); +x_68 = lean_ctor_get(x_34, 1); +lean_inc(x_68); lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_33); -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_dec(x_34); +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; } } } else { -uint8_t x_69; -lean_dec(x_28); -lean_dec(x_25); +uint8_t x_70; +lean_dec(x_29); +lean_dec(x_26); lean_dec(x_14); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_69 = !lean_is_exclusive(x_30); -if (x_69 == 0) +x_70 = !lean_is_exclusive(x_31); +if (x_70 == 0) { -return x_30; +return x_31; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_30, 0); -x_71 = lean_ctor_get(x_30, 1); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_31, 0); +x_72 = lean_ctor_get(x_31, 1); +lean_inc(x_72); lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_30); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_dec(x_31); +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; } } } else { -uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint32_t x_96; lean_object* x_97; uint32_t 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_73 = lean_ctor_get_uint8(x_10, 0); -x_74 = lean_ctor_get_uint8(x_10, 1); -x_75 = lean_ctor_get_uint8(x_10, 2); -x_76 = lean_ctor_get_uint8(x_10, 3); -x_77 = lean_ctor_get_uint8(x_10, 4); -x_78 = lean_ctor_get_uint8(x_10, 5); -x_79 = lean_ctor_get_uint8(x_10, 6); -x_80 = lean_ctor_get_uint8(x_10, 7); -x_81 = lean_ctor_get_uint8(x_10, 8); -x_82 = lean_ctor_get_uint8(x_10, 10); -x_83 = lean_ctor_get_uint8(x_10, 11); -x_84 = lean_ctor_get_uint8(x_10, 12); +uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t x_80; uint8_t x_81; uint8_t x_82; uint8_t x_83; uint8_t x_84; uint8_t x_85; uint8_t 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; uint32_t x_97; lean_object* x_98; uint32_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_74 = lean_ctor_get_uint8(x_10, 0); +x_75 = lean_ctor_get_uint8(x_10, 1); +x_76 = lean_ctor_get_uint8(x_10, 2); +x_77 = lean_ctor_get_uint8(x_10, 3); +x_78 = lean_ctor_get_uint8(x_10, 4); +x_79 = lean_ctor_get_uint8(x_10, 5); +x_80 = lean_ctor_get_uint8(x_10, 6); +x_81 = lean_ctor_get_uint8(x_10, 7); +x_82 = lean_ctor_get_uint8(x_10, 8); +x_83 = lean_ctor_get_uint8(x_10, 10); +x_84 = lean_ctor_get_uint8(x_10, 11); +x_85 = lean_ctor_get_uint8(x_10, 12); lean_dec(x_10); -x_85 = 2; -x_86 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_86, 0, x_73); -lean_ctor_set_uint8(x_86, 1, x_74); -lean_ctor_set_uint8(x_86, 2, x_75); -lean_ctor_set_uint8(x_86, 3, x_76); -lean_ctor_set_uint8(x_86, 4, x_77); -lean_ctor_set_uint8(x_86, 5, x_78); -lean_ctor_set_uint8(x_86, 6, x_79); -lean_ctor_set_uint8(x_86, 7, x_80); -lean_ctor_set_uint8(x_86, 8, x_81); -lean_ctor_set_uint8(x_86, 9, x_85); -lean_ctor_set_uint8(x_86, 10, x_82); -lean_ctor_set_uint8(x_86, 11, x_83); -lean_ctor_set_uint8(x_86, 12, x_84); -lean_ctor_set(x_4, 0, x_86); -x_87 = l_Lean_Meta_Simp_mkDefaultMethods(x_6, x_7, x_8); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +x_86 = 2; +x_87 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_87, 0, x_74); +lean_ctor_set_uint8(x_87, 1, x_75); +lean_ctor_set_uint8(x_87, 2, x_76); +lean_ctor_set_uint8(x_87, 3, x_77); +lean_ctor_set_uint8(x_87, 4, x_78); +lean_ctor_set_uint8(x_87, 5, x_79); +lean_ctor_set_uint8(x_87, 6, x_80); +lean_ctor_set_uint8(x_87, 7, x_81); +lean_ctor_set_uint8(x_87, 8, x_82); +lean_ctor_set_uint8(x_87, 9, x_86); +lean_ctor_set_uint8(x_87, 10, x_83); +lean_ctor_set_uint8(x_87, 11, x_84); +lean_ctor_set_uint8(x_87, 12, x_85); +lean_ctor_set(x_4, 0, x_87); +x_88 = l_Lean_Meta_Simp_mkDefaultMethods(x_6, x_7, x_8); +x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); -lean_dec(x_87); -x_90 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_89); -x_91 = lean_ctor_get(x_90, 0); -lean_inc(x_91); -x_92 = lean_ctor_get(x_90, 1); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); +x_91 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_90); +x_92 = lean_ctor_get(x_91, 0); lean_inc(x_92); -lean_dec(x_90); -x_93 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__3; -x_94 = lean_array_push(x_93, x_1); -x_95 = lean_box(0); -x_96 = 0; -x_97 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__1; -x_98 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; -x_99 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_94); -lean_ctor_set(x_99, 2, x_91); -lean_ctor_set(x_99, 3, x_95); -lean_ctor_set_uint32(x_99, sizeof(void*)*4, x_98); -lean_ctor_set_uint32(x_99, sizeof(void*)*4 + 4, x_96); -x_100 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__12; -x_101 = lean_st_mk_ref(x_100, x_92); -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_93 = lean_ctor_get(x_91, 1); +lean_inc(x_93); +lean_dec(x_91); +x_94 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__3; +x_95 = lean_array_push(x_94, x_1); +x_96 = lean_box(0); +x_97 = 0; +x_98 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__1; +x_99 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; +x_100 = lean_unsigned_to_nat(0u); +x_101 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_101, 0, x_98); +lean_ctor_set(x_101, 1, x_95); +lean_ctor_set(x_101, 2, x_92); +lean_ctor_set(x_101, 3, x_96); +lean_ctor_set(x_101, 4, x_100); +lean_ctor_set_uint32(x_101, sizeof(void*)*5, x_99); +lean_ctor_set_uint32(x_101, sizeof(void*)*5 + 4, x_97); +x_102 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__14; +x_103 = lean_st_mk_ref(x_102, x_93); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_102); -lean_inc(x_99); -lean_inc(x_88); -x_104 = lean_simp(x_2, x_88, x_99, x_102, x_4, x_5, x_6, x_7, x_103); -if (lean_obj_tag(x_104) == 0) +lean_inc(x_104); +lean_inc(x_101); +lean_inc(x_89); +x_106 = lean_simp(x_2, x_89, x_101, x_104, x_4, x_5, x_6, x_7, x_105); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_105; lean_object* x_106; lean_object* 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_object* x_107; lean_object* x_108; lean_object* x_109; +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); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_102); -lean_inc(x_99); -lean_inc(x_88); +lean_inc(x_104); +lean_inc(x_101); +lean_inc(x_89); lean_inc(x_3); -x_107 = lean_simp(x_3, x_88, x_99, x_102, x_4, x_5, x_6, x_7, x_106); -if (lean_obj_tag(x_107) == 0) +x_109 = lean_simp(x_3, x_89, x_101, x_104, x_4, x_5, x_6, x_7, x_108); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_ctor_get(x_105, 0); +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_110 = lean_ctor_get(x_109, 0); lean_inc(x_110); -x_111 = lean_ctor_get(x_108, 0); +x_111 = lean_ctor_get(x_109, 1); lean_inc(x_111); +lean_dec(x_109); +x_112 = lean_ctor_get(x_107, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_110, 0); +lean_inc(x_113); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_112 = l_Lean_Meta_isExprDefEq(x_110, x_111, x_4, x_5, x_6, x_7, x_109); -if (lean_obj_tag(x_112) == 0) +x_114 = l_Lean_Meta_isExprDefEq(x_112, x_113, x_4, x_5, x_6, x_7, x_111); +if (lean_obj_tag(x_114) == 0) { -lean_object* x_113; uint8_t x_114; -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_unbox(x_113); -lean_dec(x_113); -if (x_114 == 0) +lean_object* x_115; uint8_t x_116; +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_unbox(x_115); +lean_dec(x_115); +if (x_116 == 0) { -lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_108); -lean_dec(x_105); -lean_dec(x_99); -lean_dec(x_88); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +lean_dec(x_110); +lean_dec(x_107); +lean_dec(x_101); +lean_dec(x_89); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_115 = lean_ctor_get(x_112, 1); -lean_inc(x_115); -lean_dec(x_112); -x_116 = lean_st_ref_get(x_102, x_115); -lean_dec(x_102); -x_117 = lean_ctor_get(x_116, 1); +x_117 = lean_ctor_get(x_114, 1); lean_inc(x_117); -if (lean_is_exclusive(x_116)) { - lean_ctor_release(x_116, 0); - lean_ctor_release(x_116, 1); - x_118 = x_116; +lean_dec(x_114); +x_118 = lean_st_ref_get(x_104, x_117); +lean_dec(x_104); +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_116); - x_118 = lean_box(0); + lean_dec_ref(x_118); + x_120 = lean_box(0); } -if (lean_is_scalar(x_118)) { - x_119 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(0, 2, 0); } else { - x_119 = x_118; + x_121 = x_120; } -lean_ctor_set(x_119, 0, x_95); -lean_ctor_set(x_119, 1, x_117); -return x_119; +lean_ctor_set(x_121, 0, x_96); +lean_ctor_set(x_121, 1, x_119); +return x_121; } else { -lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_120 = lean_ctor_get(x_112, 1); -lean_inc(x_120); -lean_dec(x_112); -x_121 = lean_box(0); -x_122 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_108, x_105, x_121, x_88, x_99, x_102, x_4, x_5, x_6, x_7, x_120); -lean_dec(x_99); -lean_dec(x_88); -if (lean_obj_tag(x_122) == 0) +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_114, 1); +lean_inc(x_122); +lean_dec(x_114); +x_123 = lean_box(0); +x_124 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_110, x_107, x_123, x_89, x_101, x_104, x_4, x_5, x_6, x_7, x_122); +lean_dec(x_101); +lean_dec(x_89); +if (lean_obj_tag(x_124) == 0) { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_123 = lean_ctor_get(x_122, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_122, 1); -lean_inc(x_124); -lean_dec(x_122); -x_125 = lean_st_ref_get(x_102, x_124); -lean_dec(x_102); -x_126 = lean_ctor_get(x_125, 1); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_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_125)) { - lean_ctor_release(x_125, 0); - lean_ctor_release(x_125, 1); - x_127 = x_125; +lean_dec(x_124); +x_127 = lean_st_ref_get(x_104, x_126); +lean_dec(x_104); +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_125); - x_127 = lean_box(0); + lean_dec_ref(x_127); + x_129 = lean_box(0); } -if (lean_is_scalar(x_127)) { - x_128 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_129)) { + x_130 = lean_alloc_ctor(0, 2, 0); } else { - x_128 = x_127; + x_130 = x_129; } -lean_ctor_set(x_128, 0, x_123); -lean_ctor_set(x_128, 1, x_126); -return x_128; +lean_ctor_set(x_130, 0, x_125); +lean_ctor_set(x_130, 1, x_128); +return x_130; } else { -lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; -lean_dec(x_102); -x_129 = lean_ctor_get(x_122, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_122, 1); -lean_inc(x_130); -if (lean_is_exclusive(x_122)) { - lean_ctor_release(x_122, 0); - lean_ctor_release(x_122, 1); - x_131 = x_122; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +lean_dec(x_104); +x_131 = lean_ctor_get(x_124, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_124, 1); +lean_inc(x_132); +if (lean_is_exclusive(x_124)) { + lean_ctor_release(x_124, 0); + lean_ctor_release(x_124, 1); + x_133 = x_124; } else { - lean_dec_ref(x_122); - x_131 = lean_box(0); + lean_dec_ref(x_124); + x_133 = lean_box(0); } -if (lean_is_scalar(x_131)) { - x_132 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_133)) { + x_134 = lean_alloc_ctor(1, 2, 0); } else { - x_132 = x_131; + x_134 = x_133; } -lean_ctor_set(x_132, 0, x_129); -lean_ctor_set(x_132, 1, x_130); -return x_132; +lean_ctor_set(x_134, 0, x_131); +lean_ctor_set(x_134, 1, x_132); +return x_134; } } } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -lean_dec(x_108); -lean_dec(x_105); -lean_dec(x_102); -lean_dec(x_99); -lean_dec(x_88); +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_110); +lean_dec(x_107); +lean_dec(x_104); +lean_dec(x_101); +lean_dec(x_89); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_133 = lean_ctor_get(x_112, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_112, 1); -lean_inc(x_134); -if (lean_is_exclusive(x_112)) { - lean_ctor_release(x_112, 0); - lean_ctor_release(x_112, 1); - x_135 = x_112; +x_135 = lean_ctor_get(x_114, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_114, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_137 = x_114; } else { - lean_dec_ref(x_112); - x_135 = lean_box(0); + lean_dec_ref(x_114); + x_137 = lean_box(0); } -if (lean_is_scalar(x_135)) { - x_136 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); } else { - x_136 = x_135; + x_138 = x_137; } -lean_ctor_set(x_136, 0, x_133); -lean_ctor_set(x_136, 1, x_134); -return x_136; +lean_ctor_set(x_138, 0, x_135); +lean_ctor_set(x_138, 1, x_136); +return x_138; } } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_105); -lean_dec(x_102); -lean_dec(x_99); -lean_dec(x_88); +lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +lean_dec(x_107); +lean_dec(x_104); +lean_dec(x_101); +lean_dec(x_89); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_137 = lean_ctor_get(x_107, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_107, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_139 = x_107; +x_139 = lean_ctor_get(x_109, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_109, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_141 = x_109; } else { - lean_dec_ref(x_107); - x_139 = lean_box(0); + lean_dec_ref(x_109); + x_141 = lean_box(0); } -if (lean_is_scalar(x_139)) { - x_140 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); } else { - x_140 = x_139; + x_142 = x_141; } -lean_ctor_set(x_140, 0, x_137); -lean_ctor_set(x_140, 1, x_138); -return x_140; +lean_ctor_set(x_142, 0, x_139); +lean_ctor_set(x_142, 1, x_140); +return x_142; } } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -lean_dec(x_102); -lean_dec(x_99); -lean_dec(x_88); +lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; +lean_dec(x_104); +lean_dec(x_101); +lean_dec(x_89); lean_dec(x_4); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_141 = lean_ctor_get(x_104, 0); -lean_inc(x_141); -x_142 = lean_ctor_get(x_104, 1); -lean_inc(x_142); -if (lean_is_exclusive(x_104)) { - lean_ctor_release(x_104, 0); - lean_ctor_release(x_104, 1); - x_143 = x_104; +x_143 = lean_ctor_get(x_106, 0); +lean_inc(x_143); +x_144 = lean_ctor_get(x_106, 1); +lean_inc(x_144); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_145 = x_106; } else { - lean_dec_ref(x_104); - x_143 = lean_box(0); + lean_dec_ref(x_106); + x_145 = lean_box(0); } -if (lean_is_scalar(x_143)) { - x_144 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_145)) { + x_146 = lean_alloc_ctor(1, 2, 0); } else { - x_144 = x_143; + x_146 = x_145; } -lean_ctor_set(x_144, 0, x_141); -lean_ctor_set(x_144, 1, x_142); -return x_144; +lean_ctor_set(x_146, 0, x_143); +lean_ctor_set(x_146, 1, x_144); +return x_146; } } } 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; uint8_t x_151; uint8_t x_152; uint8_t x_153; uint8_t x_154; uint8_t x_155; uint8_t x_156; uint8_t 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; lean_object* x_165; uint8_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; uint32_t x_178; lean_object* x_179; uint32_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_145 = lean_ctor_get(x_4, 0); -x_146 = lean_ctor_get(x_4, 1); -x_147 = lean_ctor_get(x_4, 2); -x_148 = lean_ctor_get(x_4, 3); -x_149 = lean_ctor_get(x_4, 4); -x_150 = lean_ctor_get(x_4, 5); -x_151 = lean_ctor_get_uint8(x_4, sizeof(void*)*6); -x_152 = lean_ctor_get_uint8(x_4, sizeof(void*)*6 + 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; uint8_t x_153; uint8_t x_154; uint8_t x_155; uint8_t x_156; uint8_t 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; lean_object* x_167; uint8_t x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint32_t x_180; lean_object* x_181; uint32_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_147 = lean_ctor_get(x_4, 0); +x_148 = lean_ctor_get(x_4, 1); +x_149 = lean_ctor_get(x_4, 2); +x_150 = lean_ctor_get(x_4, 3); +x_151 = lean_ctor_get(x_4, 4); +x_152 = lean_ctor_get(x_4, 5); +x_153 = lean_ctor_get_uint8(x_4, sizeof(void*)*6); +x_154 = lean_ctor_get_uint8(x_4, sizeof(void*)*6 + 1); +lean_inc(x_152); +lean_inc(x_151); lean_inc(x_150); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_inc(x_146); -lean_inc(x_145); -lean_dec(x_4); -x_153 = lean_ctor_get_uint8(x_145, 0); -x_154 = lean_ctor_get_uint8(x_145, 1); -x_155 = lean_ctor_get_uint8(x_145, 2); -x_156 = lean_ctor_get_uint8(x_145, 3); -x_157 = lean_ctor_get_uint8(x_145, 4); -x_158 = lean_ctor_get_uint8(x_145, 5); -x_159 = lean_ctor_get_uint8(x_145, 6); -x_160 = lean_ctor_get_uint8(x_145, 7); -x_161 = lean_ctor_get_uint8(x_145, 8); -x_162 = lean_ctor_get_uint8(x_145, 10); -x_163 = lean_ctor_get_uint8(x_145, 11); -x_164 = lean_ctor_get_uint8(x_145, 12); -if (lean_is_exclusive(x_145)) { - x_165 = x_145; -} else { - lean_dec_ref(x_145); - x_165 = lean_box(0); -} -x_166 = 2; -if (lean_is_scalar(x_165)) { - x_167 = lean_alloc_ctor(0, 0, 13); -} else { - x_167 = x_165; -} -lean_ctor_set_uint8(x_167, 0, x_153); -lean_ctor_set_uint8(x_167, 1, x_154); -lean_ctor_set_uint8(x_167, 2, x_155); -lean_ctor_set_uint8(x_167, 3, x_156); -lean_ctor_set_uint8(x_167, 4, x_157); -lean_ctor_set_uint8(x_167, 5, x_158); -lean_ctor_set_uint8(x_167, 6, x_159); -lean_ctor_set_uint8(x_167, 7, x_160); -lean_ctor_set_uint8(x_167, 8, x_161); -lean_ctor_set_uint8(x_167, 9, x_166); -lean_ctor_set_uint8(x_167, 10, x_162); -lean_ctor_set_uint8(x_167, 11, x_163); -lean_ctor_set_uint8(x_167, 12, x_164); -x_168 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_168, 0, x_167); -lean_ctor_set(x_168, 1, x_146); -lean_ctor_set(x_168, 2, x_147); -lean_ctor_set(x_168, 3, x_148); -lean_ctor_set(x_168, 4, x_149); -lean_ctor_set(x_168, 5, x_150); -lean_ctor_set_uint8(x_168, sizeof(void*)*6, x_151); -lean_ctor_set_uint8(x_168, sizeof(void*)*6 + 1, x_152); -x_169 = l_Lean_Meta_Simp_mkDefaultMethods(x_6, x_7, x_8); -x_170 = lean_ctor_get(x_169, 0); -lean_inc(x_170); -x_171 = lean_ctor_get(x_169, 1); -lean_inc(x_171); -lean_dec(x_169); -x_172 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_171); -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_172, 1); -lean_inc(x_174); -lean_dec(x_172); -x_175 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__3; -x_176 = lean_array_push(x_175, x_1); -x_177 = lean_box(0); -x_178 = 0; -x_179 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__1; -x_180 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; -x_181 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_181, 0, x_179); -lean_ctor_set(x_181, 1, x_176); -lean_ctor_set(x_181, 2, x_173); -lean_ctor_set(x_181, 3, x_177); -lean_ctor_set_uint32(x_181, sizeof(void*)*4, x_180); -lean_ctor_set_uint32(x_181, sizeof(void*)*4 + 4, x_178); -x_182 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__12; -x_183 = lean_st_mk_ref(x_182, x_174); -x_184 = lean_ctor_get(x_183, 0); -lean_inc(x_184); -x_185 = lean_ctor_get(x_183, 1); -lean_inc(x_185); -lean_dec(x_183); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_168); -lean_inc(x_184); -lean_inc(x_181); -lean_inc(x_170); -x_186 = lean_simp(x_2, x_170, x_181, x_184, x_168, x_5, x_6, x_7, x_185); -if (lean_obj_tag(x_186) == 0) -{ -lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_4); +x_155 = lean_ctor_get_uint8(x_147, 0); +x_156 = lean_ctor_get_uint8(x_147, 1); +x_157 = lean_ctor_get_uint8(x_147, 2); +x_158 = lean_ctor_get_uint8(x_147, 3); +x_159 = lean_ctor_get_uint8(x_147, 4); +x_160 = lean_ctor_get_uint8(x_147, 5); +x_161 = lean_ctor_get_uint8(x_147, 6); +x_162 = lean_ctor_get_uint8(x_147, 7); +x_163 = lean_ctor_get_uint8(x_147, 8); +x_164 = lean_ctor_get_uint8(x_147, 10); +x_165 = lean_ctor_get_uint8(x_147, 11); +x_166 = lean_ctor_get_uint8(x_147, 12); +if (lean_is_exclusive(x_147)) { + x_167 = x_147; +} else { + lean_dec_ref(x_147); + x_167 = lean_box(0); +} +x_168 = 2; +if (lean_is_scalar(x_167)) { + x_169 = lean_alloc_ctor(0, 0, 13); +} else { + x_169 = x_167; +} +lean_ctor_set_uint8(x_169, 0, x_155); +lean_ctor_set_uint8(x_169, 1, x_156); +lean_ctor_set_uint8(x_169, 2, x_157); +lean_ctor_set_uint8(x_169, 3, x_158); +lean_ctor_set_uint8(x_169, 4, x_159); +lean_ctor_set_uint8(x_169, 5, x_160); +lean_ctor_set_uint8(x_169, 6, x_161); +lean_ctor_set_uint8(x_169, 7, x_162); +lean_ctor_set_uint8(x_169, 8, x_163); +lean_ctor_set_uint8(x_169, 9, x_168); +lean_ctor_set_uint8(x_169, 10, x_164); +lean_ctor_set_uint8(x_169, 11, x_165); +lean_ctor_set_uint8(x_169, 12, x_166); +x_170 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_148); +lean_ctor_set(x_170, 2, x_149); +lean_ctor_set(x_170, 3, x_150); +lean_ctor_set(x_170, 4, x_151); +lean_ctor_set(x_170, 5, x_152); +lean_ctor_set_uint8(x_170, sizeof(void*)*6, x_153); +lean_ctor_set_uint8(x_170, sizeof(void*)*6 + 1, x_154); +x_171 = l_Lean_Meta_Simp_mkDefaultMethods(x_6, x_7, x_8); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_7, x_173); +x_175 = lean_ctor_get(x_174, 0); +lean_inc(x_175); +x_176 = lean_ctor_get(x_174, 1); +lean_inc(x_176); +lean_dec(x_174); +x_177 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__3; +x_178 = lean_array_push(x_177, x_1); +x_179 = lean_box(0); +x_180 = 0; +x_181 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__1; +x_182 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__2; +x_183 = lean_unsigned_to_nat(0u); +x_184 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_184, 0, x_181); +lean_ctor_set(x_184, 1, x_178); +lean_ctor_set(x_184, 2, x_175); +lean_ctor_set(x_184, 3, x_179); +lean_ctor_set(x_184, 4, x_183); +lean_ctor_set_uint32(x_184, sizeof(void*)*5, x_182); +lean_ctor_set_uint32(x_184, sizeof(void*)*5 + 4, x_180); +x_185 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___closed__14; +x_186 = lean_st_mk_ref(x_185, x_176); x_187 = lean_ctor_get(x_186, 0); lean_inc(x_187); x_188 = lean_ctor_get(x_186, 1); @@ -1675,238 +1696,254 @@ lean_dec(x_186); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_168); -lean_inc(x_184); -lean_inc(x_181); lean_inc(x_170); -lean_inc(x_3); -x_189 = lean_simp(x_3, x_170, x_181, x_184, x_168, x_5, x_6, x_7, x_188); +lean_inc(x_187); +lean_inc(x_184); +lean_inc(x_172); +x_189 = lean_simp(x_2, x_172, x_184, x_187, x_170, x_5, x_6, x_7, x_188); if (lean_obj_tag(x_189) == 0) { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_object* x_190; lean_object* x_191; lean_object* x_192; x_190 = lean_ctor_get(x_189, 0); lean_inc(x_190); x_191 = lean_ctor_get(x_189, 1); lean_inc(x_191); lean_dec(x_189); -x_192 = lean_ctor_get(x_187, 0); -lean_inc(x_192); -x_193 = lean_ctor_get(x_190, 0); -lean_inc(x_193); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_168); -x_194 = l_Lean_Meta_isExprDefEq(x_192, x_193, x_168, x_5, x_6, x_7, x_191); -if (lean_obj_tag(x_194) == 0) +lean_inc(x_170); +lean_inc(x_187); +lean_inc(x_184); +lean_inc(x_172); +lean_inc(x_3); +x_192 = lean_simp(x_3, x_172, x_184, x_187, x_170, x_5, x_6, x_7, x_191); +if (lean_obj_tag(x_192) == 0) { -lean_object* x_195; uint8_t x_196; -x_195 = lean_ctor_get(x_194, 0); +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = lean_ctor_get(x_190, 0); lean_inc(x_195); -x_196 = lean_unbox(x_195); -lean_dec(x_195); -if (x_196 == 0) +x_196 = lean_ctor_get(x_193, 0); +lean_inc(x_196); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_170); +x_197 = l_Lean_Meta_isExprDefEq(x_195, x_196, x_170, x_5, x_6, x_7, x_194); +if (lean_obj_tag(x_197) == 0) +{ +lean_object* x_198; uint8_t x_199; +x_198 = lean_ctor_get(x_197, 0); +lean_inc(x_198); +x_199 = lean_unbox(x_198); +lean_dec(x_198); +if (x_199 == 0) { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +lean_dec(x_193); lean_dec(x_190); -lean_dec(x_187); -lean_dec(x_181); +lean_dec(x_184); +lean_dec(x_172); lean_dec(x_170); -lean_dec(x_168); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_197 = lean_ctor_get(x_194, 1); -lean_inc(x_197); -lean_dec(x_194); -x_198 = lean_st_ref_get(x_184, x_197); -lean_dec(x_184); -x_199 = lean_ctor_get(x_198, 1); -lean_inc(x_199); -if (lean_is_exclusive(x_198)) { - lean_ctor_release(x_198, 0); - lean_ctor_release(x_198, 1); - x_200 = x_198; +x_200 = lean_ctor_get(x_197, 1); +lean_inc(x_200); +lean_dec(x_197); +x_201 = lean_st_ref_get(x_187, x_200); +lean_dec(x_187); +x_202 = lean_ctor_get(x_201, 1); +lean_inc(x_202); +if (lean_is_exclusive(x_201)) { + lean_ctor_release(x_201, 0); + lean_ctor_release(x_201, 1); + x_203 = x_201; } else { - lean_dec_ref(x_198); - x_200 = lean_box(0); + lean_dec_ref(x_201); + x_203 = lean_box(0); } -if (lean_is_scalar(x_200)) { - x_201 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_203)) { + x_204 = lean_alloc_ctor(0, 2, 0); } else { - x_201 = x_200; + x_204 = x_203; } -lean_ctor_set(x_201, 0, x_177); -lean_ctor_set(x_201, 1, x_199); -return x_201; +lean_ctor_set(x_204, 0, x_179); +lean_ctor_set(x_204, 1, x_202); +return x_204; } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; -x_202 = lean_ctor_get(x_194, 1); -lean_inc(x_202); -lean_dec(x_194); -x_203 = lean_box(0); -x_204 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_190, x_187, x_203, x_170, x_181, x_184, x_168, x_5, x_6, x_7, x_202); -lean_dec(x_181); -lean_dec(x_170); -if (lean_obj_tag(x_204) == 0) -{ -lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_205 = lean_ctor_get(x_204, 0); +lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_205 = lean_ctor_get(x_197, 1); lean_inc(x_205); -x_206 = lean_ctor_get(x_204, 1); -lean_inc(x_206); -lean_dec(x_204); -x_207 = lean_st_ref_get(x_184, x_206); +lean_dec(x_197); +x_206 = lean_box(0); +x_207 = l_Lean_Elab_Tactic_NormCast_proveEqUsing___lambda__1(x_3, x_193, x_190, x_206, x_172, x_184, x_187, x_170, x_5, x_6, x_7, x_205); lean_dec(x_184); -x_208 = lean_ctor_get(x_207, 1); +lean_dec(x_172); +if (lean_obj_tag(x_207) == 0) +{ +lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_208 = lean_ctor_get(x_207, 0); lean_inc(x_208); -if (lean_is_exclusive(x_207)) { - lean_ctor_release(x_207, 0); - lean_ctor_release(x_207, 1); - x_209 = x_207; +x_209 = lean_ctor_get(x_207, 1); +lean_inc(x_209); +lean_dec(x_207); +x_210 = lean_st_ref_get(x_187, x_209); +lean_dec(x_187); +x_211 = lean_ctor_get(x_210, 1); +lean_inc(x_211); +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + x_212 = x_210; } else { - lean_dec_ref(x_207); - x_209 = lean_box(0); + lean_dec_ref(x_210); + x_212 = lean_box(0); } -if (lean_is_scalar(x_209)) { - x_210 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_212)) { + x_213 = lean_alloc_ctor(0, 2, 0); } else { - x_210 = x_209; + x_213 = x_212; } -lean_ctor_set(x_210, 0, x_205); -lean_ctor_set(x_210, 1, x_208); -return x_210; +lean_ctor_set(x_213, 0, x_208); +lean_ctor_set(x_213, 1, x_211); +return x_213; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_184); -x_211 = lean_ctor_get(x_204, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_204, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_204)) { - lean_ctor_release(x_204, 0); - lean_ctor_release(x_204, 1); - x_213 = x_204; +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +lean_dec(x_187); +x_214 = lean_ctor_get(x_207, 0); +lean_inc(x_214); +x_215 = lean_ctor_get(x_207, 1); +lean_inc(x_215); +if (lean_is_exclusive(x_207)) { + lean_ctor_release(x_207, 0); + lean_ctor_release(x_207, 1); + x_216 = x_207; } else { - lean_dec_ref(x_204); - x_213 = lean_box(0); + lean_dec_ref(x_207); + x_216 = lean_box(0); } -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_216)) { + x_217 = lean_alloc_ctor(1, 2, 0); } else { - x_214 = x_213; + x_217 = x_216; } -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set(x_214, 1, x_212); -return x_214; +lean_ctor_set(x_217, 0, x_214); +lean_ctor_set(x_217, 1, x_215); +return x_217; } } } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; +lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; +lean_dec(x_193); lean_dec(x_190); lean_dec(x_187); lean_dec(x_184); -lean_dec(x_181); +lean_dec(x_172); lean_dec(x_170); -lean_dec(x_168); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_215 = lean_ctor_get(x_194, 0); -lean_inc(x_215); -x_216 = lean_ctor_get(x_194, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_217 = x_194; +x_218 = lean_ctor_get(x_197, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_197, 1); +lean_inc(x_219); +if (lean_is_exclusive(x_197)) { + lean_ctor_release(x_197, 0); + lean_ctor_release(x_197, 1); + x_220 = x_197; } else { - lean_dec_ref(x_194); - x_217 = lean_box(0); + lean_dec_ref(x_197); + x_220 = lean_box(0); } -if (lean_is_scalar(x_217)) { - x_218 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_220)) { + x_221 = lean_alloc_ctor(1, 2, 0); } else { - x_218 = x_217; + x_221 = x_220; } -lean_ctor_set(x_218, 0, x_215); -lean_ctor_set(x_218, 1, x_216); -return x_218; +lean_ctor_set(x_221, 0, x_218); +lean_ctor_set(x_221, 1, x_219); +return x_221; } } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; +lean_dec(x_190); lean_dec(x_187); lean_dec(x_184); -lean_dec(x_181); +lean_dec(x_172); lean_dec(x_170); -lean_dec(x_168); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_219 = lean_ctor_get(x_189, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_189, 1); -lean_inc(x_220); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_221 = x_189; +x_222 = lean_ctor_get(x_192, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_192, 1); +lean_inc(x_223); +if (lean_is_exclusive(x_192)) { + lean_ctor_release(x_192, 0); + lean_ctor_release(x_192, 1); + x_224 = x_192; } else { - lean_dec_ref(x_189); - x_221 = lean_box(0); + lean_dec_ref(x_192); + x_224 = lean_box(0); } -if (lean_is_scalar(x_221)) { - x_222 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_224)) { + x_225 = lean_alloc_ctor(1, 2, 0); } else { - x_222 = x_221; + x_225 = x_224; } -lean_ctor_set(x_222, 0, x_219); -lean_ctor_set(x_222, 1, x_220); -return x_222; +lean_ctor_set(x_225, 0, x_222); +lean_ctor_set(x_225, 1, x_223); +return x_225; } } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; +lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +lean_dec(x_187); lean_dec(x_184); -lean_dec(x_181); +lean_dec(x_172); lean_dec(x_170); -lean_dec(x_168); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_223 = lean_ctor_get(x_186, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_186, 1); -lean_inc(x_224); -if (lean_is_exclusive(x_186)) { - lean_ctor_release(x_186, 0); - lean_ctor_release(x_186, 1); - x_225 = x_186; +x_226 = lean_ctor_get(x_189, 0); +lean_inc(x_226); +x_227 = lean_ctor_get(x_189, 1); +lean_inc(x_227); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_228 = x_189; } else { - lean_dec_ref(x_186); - x_225 = lean_box(0); + lean_dec_ref(x_189); + x_228 = lean_box(0); } -if (lean_is_scalar(x_225)) { - x_226 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 2, 0); } else { - x_226 = x_225; + x_229 = x_228; } -lean_ctor_set(x_226, 0, x_223); -lean_ctor_set(x_226, 1, x_224); -return x_226; +lean_ctor_set(x_229, 0, x_226); +lean_ctor_set(x_229, 1, x_227); +return x_229; } } } @@ -1923,119 +1960,7 @@ lean_dec(x_4); return x_13; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; double x_18; double x_19; double x_20; double x_21; double x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_15 = lean_ctor_get(x_13, 0); -x_16 = 0; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Float_ofScientific(x_8, x_16, x_17); -lean_dec(x_8); -x_19 = l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2___closed__1; -x_20 = lean_float_div(x_18, x_19); -x_21 = l_Float_ofScientific(x_15, x_16, x_17); -lean_dec(x_15); -x_22 = lean_float_div(x_21, x_19); -x_23 = lean_box_float(x_20); -x_24 = lean_box_float(x_22); -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, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_13, 0, x_26); -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_dec(x_13); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_8, x_29, x_30); -lean_dec(x_8); -x_32 = l_Lean_withStartStopSeconds___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = l_Float_ofScientific(x_27, x_29, x_30); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_32); -x_36 = lean_box_float(x_33); -x_37 = lean_box_float(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_11); -lean_ctor_set(x_39, 1, 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_28); -return x_40; -} -} -else -{ -uint8_t x_41; -lean_dec(x_8); -x_41 = !lean_is_exclusive(x_10); -if (x_41 == 0) -{ -return x_10; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_10, 0); -x_43 = lean_ctor_get(x_10, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_10); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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: { if (lean_obj_tag(x_1) == 0) @@ -2060,89 +1985,23 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_6); lean_inc(x_9); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__3(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__2(x_4, x_7, x_8, x_9, x_10, x_13); lean_dec(x_9); lean_dec(x_4); return x_14; } } -static double _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -2152,7 +2011,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -2160,11 +2019,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___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, uint8_t x_8, double x_9, double 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: { double x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__1; +x_17 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_18 = lean_alloc_ctor(0, 2, 17); @@ -2173,7 +2032,7 @@ lean_ctor_set(x_18, 1, x_3); lean_ctor_set_float(x_18, sizeof(void*)*2, x_17); lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_17); lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_2); -x_19 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__2; +x_19 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2___closed__2; x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_19); if (x_20 == 0) { @@ -2183,7 +2042,7 @@ lean_object* x_21; lean_object* x_22; lean_dec(x_3); lean_dec(x_1); x_21 = lean_box(0); -x_22 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); return x_22; } else @@ -2197,7 +2056,7 @@ lean_ctor_set_float(x_23, sizeof(void*)*2, x_9); lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_2); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); +x_25 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); return x_25; } } @@ -2212,12 +2071,12 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); +x_28 = l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); return x_28; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Elab_Tactic_NormCast_proveEqUsingDown___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -2225,16 +2084,16 @@ x_1 = lean_mk_string_from_bytes(" #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -58,6 +58,7 @@ 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*); +lean_object* initialize_Lean_Meta_Diagnostics(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; @@ -198,6 +199,9 @@ 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); +res = initialize_Lean_Meta_Diagnostics(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/Basic.c b/stage0/stdlib/Lean/Meta/Basic.c index 3d02be6b7629..088a7cf46a16 100644 --- a/stage0/stdlib/Lean/Meta/Basic.c +++ b/stage0/stdlib/Lean/Meta/Basic.c @@ -49,6 +49,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalConte LEAN_EXPORT lean_object* l_Lean_Meta_Context_synthPendingDepth___default; LEAN_EXPORT lean_object* l_Lean_Meta_withReplaceFVarId(lean_object*); static lean_object* l_Lean_Meta_mkFunUnit___closed__4; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2___closed__1; lean_object* l_Lean_MetavarContext_findDecl_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instAlternativeMetaM___closed__3; LEAN_EXPORT lean_object* l_Lean_MVarId_withContext(lean_object*); @@ -61,6 +62,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___rarg___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___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__5; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -69,10 +71,10 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalIn 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_instMonadEnvMetaM___closed__2; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_is_level_def_eq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkFunUnit___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -size_t lean_usize_shift_right(size_t, size_t); lean_object* l_Lean_Level_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_LMVarId_getLevel___closed__2; @@ -83,8 +85,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_synthPending___boxed(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Meta_processPostponed_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshLevelMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapError___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_orelseMergeErrors(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_getDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(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_MVarId_withContext___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Context_config___default; LEAN_EXPORT lean_object* l_Lean_Meta_mkForallFVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,12 +96,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getMVarDeclKind___boxed(lean_object*, lean_ LEAN_EXPORT 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*); static lean_object* l_Lean_Meta_ParamInfo_backDeps___default___closed__1; static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop(lean_object*); -extern lean_object* l_Lean_diag; LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar___spec__1___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_whnfI(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_abstract(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_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*); @@ -114,7 +117,6 @@ lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, l LEAN_EXPORT uint8_t l_Lean_Meta_Config_assignSyntheticOpaque___default; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__11; extern lean_object* l_Lean_maxRecDepthErrorMessage; -uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_withConfig(lean_object*); lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getTheoremInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -128,16 +130,16 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_getLocalDeclFromUserNa uint64_t lean_uint64_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instBEqInfoCacheKey; uint64_t lean_uint64_mix_hash(uint64_t, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___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_throwError___at_Lean_Meta_collectForwardDeps___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Context_canUnfold_x3f___default; lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImp(lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__8; -size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_withAtLeastTransparency___rarg___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_withLocalDecls_loop___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaTelescope(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -148,11 +150,8 @@ static lean_object* l_Lean_Meta_getLocalDeclFromUserName___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608_(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Diagnostics_unfoldCounter___default; lean_object* l_Lean_Level_succ___override(lean_object*); -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_withLocalInstancesImp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isProp___default; @@ -168,12 +167,10 @@ lean_object* l_Lean_mkLevelMax_x27(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_transparency___default; LEAN_EXPORT lean_object* l_Lean_Meta_withAssignableSyntheticOpaque___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__11; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Meta_mkFreshLevelMVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope(lean_object*); 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_instAlternativeMetaM___closed__1; LEAN_EXPORT lean_object* l_Nat_foldM_loop___at_Lean_Meta_mkFreshLevelMVars___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_Meta_Basic_0__Lean_Meta_isClassQuickConst_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -181,19 +178,16 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spe static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux___closed__2; static lean_object* l_Lean_Meta_liftMkBindingM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_FVarId_throwUnknown(lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__6; uint8_t lean_usize_dec_eq(size_t, size_t); static lean_object* l_Lean_Meta_instMonadMetaM___closed__5; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_isLetVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withTrackingZetaDelta___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkpointDefEq___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarCore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instAlternativeMetaM___lambda__1___closed__2; static lean_object* l_Lean_Meta_instInhabitedMetaM___rarg___closed__1; @@ -208,7 +202,6 @@ static lean_object* l_Lean_Meta_mkFunUnit___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_appendOptMessageData(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstance___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t); @@ -220,7 +213,6 @@ LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___at___private_Lean_Meta_B LEAN_EXPORT lean_object* l_Lean_Meta_etaExpand___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_normalizeLevel(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_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* l_Lean_Expr_bvar___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshLevelMVarsFor___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -229,13 +221,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTel LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCache_instances___default; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstances___rarg___lambda__1___boxed(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*); static lean_object* l_Lean_Meta_instMonadEnvMetaM___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_State_zetaDeltaFVarIds___default; -lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); static lean_object* l_Lean_Meta_State_postponed___default___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqNoConstantApprox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withAssignableSyntheticOpaque(lean_object*); @@ -251,14 +240,12 @@ lean_object* l_Lean_Expr_fvarId_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withInstImplicitAsImplict(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_getKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cache_defEqPerm___default; -static lean_object* l_Lean_Meta_appendOptMessageData___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_isLevelDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaMetaTelescope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_instInhabitedPostponedEntry___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstanceImp(lean_object*); static lean_object* l_Lean_Meta_instInhabitedCache___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -269,9 +256,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withConfig___rarg___lambda__1(lean_object*, LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9; static lean_object* l_Lean_MVarId_getDecl___closed__2; +static lean_object* l_Lean_Meta_MetaM_toIO___rarg___closed__2; LEAN_EXPORT 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___private_Lean_Meta_Basic_0__Lean_Meta_withLocalContextImp(lean_object*); -lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg(lean_object*); extern lean_object* l_Lean_Compiler_inlineAttrs; @@ -291,7 +278,7 @@ lean_object* lean_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarAtCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Expr_cleanupAnnotations(lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__3; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__15; lean_object* l_Lean_EnumAttributes_setValue___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope(lean_object*); @@ -332,16 +319,18 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate___lambda__1___boxed(le LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Cache_inferType___default; static lean_object* l_Lean_Meta_instInhabitedCache___closed__4; +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclD(lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_findDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_getKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__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_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getNumPostponed___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_higherOrderOutParam___default; LEAN_EXPORT lean_object* l_Lean_Meta_State_cache___default; lean_object* l_Lean_LocalDecl_index(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2(lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__2; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_State_postponed___default; LEAN_EXPORT lean_object* l_Lean_Meta_findLocalDecl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -349,7 +338,6 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_getDecl___spec__1___ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_5____closed__1; static lean_object* l_Lean_Meta_mkLevelStuckErrorMessage___closed__1; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1; -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getZetaDeltaFVarIds(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); @@ -367,7 +355,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstance___at___private_Lean_Me LEAN_EXPORT lean_object* l_Lean_Meta_whnfAtMostI(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getNumPostponed___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_Meta_isExprDefEq___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__1; size_t lean_usize_of_nat(lean_object*); @@ -389,7 +376,6 @@ LEAN_EXPORT lean_object* l_Lean_isReducible___at___private_Lean_Meta_Basic_0__Le LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp_process___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedDefEqCache; LEAN_EXPORT lean_object* l_Lean_FVarId_getDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withReducible___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -417,12 +403,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_FunInfo_resultDeps___default; LEAN_EXPORT lean_object* l_Lean_Meta_fullApproxDefEq___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_constApprox___default; uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_reportDiag___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_univApprox___default; LEAN_EXPORT lean_object* l_Lean_Expr_abstractRangeM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_funInfo___default___closed__2; -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_getConstTemp_x3f___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCache_default___default; @@ -431,14 +415,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Cache_whnfAll___default; LEAN_EXPORT uint8_t l_Lean_Meta_Config_unificationHints___default; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarWithIdCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* l_Lean_ppExprWithInfos(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___rarg___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_Meta_withLocalInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_mkMessageBodyFor_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_synthInstance___default___closed__6; -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__5; LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isImplicit(lean_object*); static lean_object* l_Lean_Meta_instMonadMCtxMetaM___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___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*); @@ -449,13 +430,14 @@ static lean_object* l_Lean_Meta_Cache_synthInstance___default___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_numLevelParams(lean_object*); -static lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___closed__1; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLevelStuckErrorMessage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setPostponed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_orElse___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(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*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_withContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -486,18 +468,14 @@ LEAN_EXPORT lean_object* l_Lean_Meta_liftMkBindingM___rarg___boxed(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*); -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_reportDiag___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedPostponedEntry; -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6; uint8_t l_Lean_Expr_isMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_isReducible___at___private_Lean_Meta_Basic_0__Lean_Meta_getDefInfoTemp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAux___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_Meta_collectAboveThreshold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -525,14 +503,12 @@ LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_Meta_setInlineAttribute___spec_ static lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_withLocalDeclsD___spec__1(lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_instAlternativeMetaM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM___closed__1; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_elimMVarDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_setMVarKind(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withMVarContextImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -542,11 +518,12 @@ static lean_object* l_Lean_FVarId_throwUnknown___rarg___closed__1; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshTypeMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwUnknownFVar___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___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_Meta_collectForwardDeps(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarWithIdCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_addExprMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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_Lean_Meta_withLocalInstances(lean_object*); static lean_object* l_Lean_Meta_instAddMessageContextMetaM___closed__1; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1___closed__2; @@ -554,14 +531,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withDefault___rarg(lean_object*, lean_objec static lean_object* l_Lean_Meta_isInductivePredicate___closed__1; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_350____spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_liftMkBindingM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls_loop___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_inferType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_MessageData_nil; LEAN_EXPORT lean_object* l_Lean_Meta_withTrackingZetaDelta___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -576,7 +552,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescop static lean_object* l_Lean_Meta_instAlternativeMetaM___closed__2; LEAN_EXPORT lean_object* l_Lean_MVarId_setType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_hasMVar(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalInstances_erase(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -594,20 +569,19 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_MVarId_getDecl___spec__1(le LEAN_EXPORT lean_object* l_Lean_Meta_Cache_whnfDefault___default; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_liftMkBindingM___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__7; LEAN_EXPORT lean_object* l_Lean_FVarId_getValue_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_shouldReduceAll(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_350____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapError___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM___closed__3; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); -static double l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLetFVars(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getParamNames___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFunUnit(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLetFVars___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_instMonadMCtxMetaM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_abstract___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -616,13 +590,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___rarg(lean_object* static lean_object* l_Lean_Meta_Cache_synthInstance___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_ParamInfo_isStrictImplicit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_saveState___boxed(lean_object*); -extern lean_object* l_Lean_warningAsError; LEAN_EXPORT lean_object* l_Lean_Meta_whnf___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___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_Meta_mapMetaM(lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__17; -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_MkBinding_collectForwardDeps(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -633,11 +604,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance___rarg___lambda__1( LEAN_EXPORT lean_object* l_Lean_Meta_withoutProofIrrelevance___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getParamNames___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Diagnostics_heuristicCounter___default; +extern lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_MetavarContext_getExprAssignmentCore_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isReadOnlyOrSyntheticOpaqueExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setMVarKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCache_all___default; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SavedState_restore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_State_diag___default; LEAN_EXPORT lean_object* l_Lean_Meta_getResetPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -648,17 +619,17 @@ LEAN_EXPORT lean_object* l_Lean_Meta_resetZetaDeltaFVarIds___rarg___boxed(lean_o LEAN_EXPORT lean_object* l_Lean_FVarId_getBinderInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalDecl_isLet(lean_object*); lean_object* l_Lean_MetavarContext_mkBinding(uint8_t, lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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*); static lean_object* l_Lean_Meta_DefEqCache_reducible___default___closed__5; LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isInstImplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_liftMkBindingM___spec__1(lean_object*); static lean_object* l_Lean_Meta_State_mctx___default___closed__2; LEAN_EXPORT 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_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_saveState___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__14; +extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ParamInfo_isImplicit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -668,6 +639,7 @@ LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Meta_Ba lean_object* l_Lean_LocalDecl_fvarId(lean_object*); static lean_object* l_Lean_Meta_mkFunUnit___closed__2; lean_object* l_Lean_MetavarContext_MkBinding_elimMVarDeps(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__5___closed__3; static lean_object* l_Lean_Meta_processPostponed_loop___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_ppExpr___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -684,34 +656,30 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_liftMkBindingM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassExpensive_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instMonadBacktrackSavedStateMetaM; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DefEqCache_reducible___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_getZetaDeltaFVarIds___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_approxDefEq___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReplaceFVarId___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -extern lean_object* l_Lean_diag_threshold; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_withInstImplicitAsImplict___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getParamNames(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mod(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_usize_to_nat(size_t); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_LMVarId_getLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_getMVarDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedMetaM___rarg___closed__2; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); 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*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__6; 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*); extern lean_object* l_Lean_inheritedTraceOptions; -static lean_object* l_Lean_Meta_appendOptMessageData___closed__2; lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_orelseMergeErrors___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2(lean_object*); @@ -726,7 +694,6 @@ uint8_t l_Lean_LocalDecl_isImplementationDetail(lean_object*); static lean_object* l_Lean_Meta_Cache_funInfo___default___closed__1; static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_abstractRange(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_instAlternativeMetaM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MVarId_getDecl___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instantiateLambda___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -743,7 +710,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaTelescopeReducing___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaBoundedTelescope(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkLevelErrorMessage(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewBinderInfosImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__1___boxed(lean_object*, lean_object*); @@ -756,11 +722,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SavedState_restoreFull___boxed(lean_object* LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_withLocalDecls_loop___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducibleAndInstances___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedCache; lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getLevelMVarAssignment_x3f___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -770,7 +734,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp_process___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Cache_inferType___default___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFunUnit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstances(lean_object*); @@ -781,25 +744,19 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescop LEAN_EXPORT lean_object* l_Lean_Meta_throwUnknownFVar___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Config_etaStruct___default; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1; -static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_setInlineAttribute___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM; LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getPostponed___boxed(lean_object*); -static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_FunInfo_paramInfo___default; static lean_object* l_Lean_Meta_instAlternativeMetaM___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_recordUnfold___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isStrictImplicit(lean_object*); -LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_fvarsSizeLtMaxFVars(lean_object*, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_Cache_new(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_recordInstance___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_isDecInst___default; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__4; LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withMVarContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_State_mctx___default___closed__1; @@ -807,7 +764,6 @@ static lean_object* l_Lean_Meta_setInlineAttribute___closed__1; static lean_object* l_Lean_Meta_checkpointDefEq___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_mapErrorImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___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_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescope___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewLocalInstance___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getResetPostponed___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -819,6 +775,7 @@ static lean_object* l_Lean_Meta_instMonadMCtxMetaM___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshLevelMVarsFor(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_getFVarLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_fullApproxDefEq(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVarWithId(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -826,7 +783,6 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_sortFVarIds___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshExprMVar(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mapForallTelescope_x27___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_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -836,7 +792,6 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_ lean_object* l_Lean_MessageData_ofLevel(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Context_localInstances___default; LEAN_EXPORT lean_object* l_Lean_Expr_abstractRangeM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -844,27 +799,23 @@ LEAN_EXPORT uint8_t l_Lean_Meta_Context_univApprox___default; LEAN_EXPORT lean_object* l_Lean_Meta_savingCache(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp(lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); +static double l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__4; lean_object* l_ReaderT_instFunctorOfMonad___rarg(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_instInhabitedCache___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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_Meta_withLocalInstancesImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_orElse(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency___rarg___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_Context_config___default___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_fvarsSizeLtMaxFVars___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_getType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_liftMetaM___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_setType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_instAlternativeMetaM___spec__1(lean_object*); @@ -876,12 +827,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_collectForwardDeps___boxed(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Meta_setMVarType___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1___closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__2; static lean_object* l_Lean_FVarId_throwUnknown___rarg___closed__2; static double l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_run(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_savingCache___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__2; -uint64_t l_Lean_Name_hash___override(lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_withIncRecDepth___spec__1___rarg___closed__2; static lean_object* l_Lean_Meta_instMonadMetaM___closed__1; LEAN_EXPORT lean_object* l_Lean_mkFreshFVarId___at___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux_process___spec__1___rarg(lean_object*, lean_object*); @@ -896,11 +847,10 @@ lean_object* l_Lean_PersistentHashMap_insert___at_Lean_assignLevelMVar___spec__1 static lean_object* l_Lean_Meta_instInhabitedDefEqCache___closed__1; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; extern lean_object* l_Lean_Expr_instHashable; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDeclsD___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadEnvMetaM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLetDeclImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Basic_0__Lean_Meta_getConstTemp_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Init_MetaTypes_0__Lean_Meta_beqTransparencyMode____x40_Init_MetaTypes___hyg_67_(uint8_t, uint8_t); @@ -912,7 +862,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaTelescopeReducing(lean_object*, l lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_ReplaceImpl_replaceUnsafeM_visit___at___private_Lean_Meta_Basic_0__Lean_Meta_exposeRelevantUniverses___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Context_lctx___default___closed__2; -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); LEAN_EXPORT lean_object* l_Lean_assignLevelMVar___at_Lean_Meta_normalizeLevel___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -925,14 +874,12 @@ static lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAuxAux___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_Meta_map2MetaM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_350____boxed(lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_shouldReduceAll___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_elimMVarDeps(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Context_inTypeClassResolution___default; LEAN_EXPORT lean_object* l_Lean_Meta_instantiateMVarsIfMVarApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withAtLeastTransparency___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__12; LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_getParamNames___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -946,20 +893,15 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at_Lean_Meta_mkFreshExprMVarAt___spe LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkLevelIMax_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewLocalInstancesImpAux___at_Lean_Meta_withNewLocalInstances___spec__1(lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__7; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__9; static lean_object* l_Lean_Meta_instInhabitedParamInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_isReadOnlyExprMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instAlternativeMetaM; -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; LEAN_EXPORT lean_object* l_Lean_Meta_ParamInfo_isInstImplicit___boxed(lean_object*); -size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalInstancesImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_processPostponed___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_16608____closed__4; LEAN_EXPORT 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_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope(lean_object*); @@ -982,13 +924,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMCtxMetaM___lambda__2(lean_object* static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_etaExpand___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_instMonadMCtxMetaM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withExistingLocalDecls(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLevelMVarDepth___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_approxDefEqImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux(lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_Basic_0__Lean_Meta_beqInfoCacheKey____x40_Lean_Meta_Basic___hyg_350_(lean_object*, lean_object*); uint8_t l_Lean_Level_any(lean_object*, lean_object*); static lean_object* l_Lean_Meta_processPostponed___lambda__1___closed__4; @@ -1012,18 +952,16 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshLMVarId___at_Lean_Meta_mkFreshLevelMVar__ lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux_process(uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_ParamInfo_binderInfo___default; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -size_t lean_usize_shift_left(size_t, size_t); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___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_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_recordInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withIncRecDepth___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingAux___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaLetTelescope___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_Meta_setInlineAttribute(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__12; lean_object* l_Lean_MetavarContext_setMVarKind(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_sortFVarIds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1035,6 +973,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_withInstImplic LEAN_EXPORT lean_object* l_Lean_Meta_mkFreshTypeMVar(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withTransparency___rarg___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getTheoremInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_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*); @@ -1042,9 +981,10 @@ LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM(lean_object*, lean_object*, lean_ 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*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Meta_Basic_0__Lean_Meta_processPostponedStep___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SavedState_restoreFull(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getPostponed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassQuick_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1058,9 +998,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___rarg___lambda__1___boxed(le LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_isClassImp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); LEAN_EXPORT lean_object* l_Lean_Meta_processPostponed_loop___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalDeclFromUserName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1068,7 +1006,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_modifyPostponed___boxed(lean_object*, lean_ static lean_object* l_Lean_Meta_Cache_synthInstance___default___closed__3; static lean_object* l_Lean_Meta_liftMkBindingM___rarg___closed__2; lean_object* l_Lean_simpLevelMax_x27(lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); +static lean_object* l_Lean_initFn____x40_Lean_Meta_Basic___hyg_15843____closed__4; lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_level_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducible___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1088,7 +1026,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isDefEq(lean_object*, lean_object*, lean_ob static size_t l_Lean_Meta_instInhabitedState___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_orelseMergeErrorsImp___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_getMVarDeclKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instMonadMCtxMetaM; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateLambdaAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkFunUnit___closed__1; @@ -1101,7 +1038,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getFVarLocalDecl___boxed(lean_object*, lean static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_instantiateLambda(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_FVarId_throwUnknown___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___rarg___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_Exception_isRuntime(lean_object*); @@ -1122,7 +1058,6 @@ lean_object* l_Lean_MetavarContext_incDepth(lean_object*, uint8_t); uint8_t l_Lean_Expr_isForall(lean_object*); static lean_object* l_Lean_Meta_State_postponed___default___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_reportDiag___closed__1; LEAN_EXPORT lean_object* l_Lean_MVarId_setKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withMVarContext(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1142,7 +1077,6 @@ 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*); @@ -1152,13 +1086,10 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_Cache_store(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__5(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*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getPostponed___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l_Lean_Name_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isReadOnlyLevelMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1170,6 +1101,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescop static lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_findMVarDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_saveState(lean_object*); +static double l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__5; static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4; lean_object* lean_expr_abstract_range(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withReducible(lean_object*); @@ -1179,11 +1111,11 @@ LEAN_EXPORT lean_object* l_Lean_FVarId_getValue_x3f(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_Meta_getLevelMVarDepth(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); lean_object* l_Lean_LocalContext_replaceFVarId(lean_object*, lean_object*, lean_object*); -size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_whnfR(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkFreshMVarId___at_Lean_Meta_mkFreshExprMVarAt___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withDefault(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallMetaBoundedTelescope___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3645,49 +3577,42 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_MetaM_run_x27___rarg), 6, 0); return x_2; } } -static lean_object* _init_l_Lean_Meta_MetaM_toIO___rarg___closed__1() { +static lean_object* _init_l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("internal exception #", 20); +x_1 = l_Lean_maxRecDepth; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___rarg___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_io_get_num_heartbeats(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = !lean_is_exclusive(x_2); +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; 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_2, 8); +x_11 = lean_ctor_get(x_7, 4); lean_dec(x_11); -lean_ctor_set(x_2, 8, x_8); -x_12 = lean_st_mk_ref(x_3, 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); +x_12 = lean_ctor_get(x_7, 2); lean_dec(x_12); -x_15 = lean_st_mk_ref(x_5, x_14); +x_13 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1; +x_14 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_13); +lean_ctor_set(x_7, 4, x_14); +lean_ctor_set(x_7, 2, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*11 + 1, x_2); +x_15 = lean_st_mk_ref(x_3, 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); -lean_inc(x_13); lean_inc(x_16); -x_18 = lean_apply_5(x_1, x_4, x_16, x_2, x_13, x_17); +x_18 = lean_apply_5(x_4, x_5, x_16, x_7, x_8, x_17); if (lean_obj_tag(x_18) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); @@ -3695,6088 +3620,4687 @@ lean_inc(x_20); lean_dec(x_18); x_21 = lean_st_ref_get(x_16, x_20); lean_dec(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_st_ref_get(x_13, x_23); -lean_dec(x_13); -x_25 = !lean_is_exclusive(x_24); -if (x_25 == 0) +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_24, 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_19); +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; 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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 0, x_19); +lean_ctor_set(x_27, 1, x_25); x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_19); -lean_ctor_set(x_28, 1, x_27); -lean_ctor_set(x_24, 0, x_28); -return x_24; +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_26); +return x_28; +} +} +else +{ +uint8_t x_29; +lean_dec(x_16); +x_29 = !lean_is_exclusive(x_18); +if (x_29 == 0) +{ +return x_18; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_29 = lean_ctor_get(x_24, 0); -x_30 = lean_ctor_get(x_24, 1); +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_18, 0); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_24); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_22); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_19); +lean_dec(x_18); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_30); -return x_33; +return x_32; +} } } else { -lean_object* x_34; -lean_dec(x_16); -lean_dec(x_13); -x_34 = lean_ctor_get(x_18, 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; 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; +x_33 = lean_ctor_get(x_7, 0); +x_34 = lean_ctor_get(x_7, 1); +x_35 = lean_ctor_get(x_7, 3); +x_36 = lean_ctor_get(x_7, 5); +x_37 = lean_ctor_get(x_7, 6); +x_38 = lean_ctor_get(x_7, 7); +x_39 = lean_ctor_get(x_7, 8); +x_40 = lean_ctor_get(x_7, 9); +x_41 = lean_ctor_get(x_7, 10); +x_42 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +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_inc(x_34); -if (lean_obj_tag(x_34) == 0) +lean_inc(x_33); +lean_dec(x_7); +x_43 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1; +x_44 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_43); +x_45 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_45, 0, x_33); +lean_ctor_set(x_45, 1, x_34); +lean_ctor_set(x_45, 2, x_1); +lean_ctor_set(x_45, 3, x_35); +lean_ctor_set(x_45, 4, x_44); +lean_ctor_set(x_45, 5, x_36); +lean_ctor_set(x_45, 6, x_37); +lean_ctor_set(x_45, 7, x_38); +lean_ctor_set(x_45, 8, x_39); +lean_ctor_set(x_45, 9, x_40); +lean_ctor_set(x_45, 10, x_41); +lean_ctor_set_uint8(x_45, sizeof(void*)*11, x_42); +lean_ctor_set_uint8(x_45, sizeof(void*)*11 + 1, x_2); +x_46 = lean_st_mk_ref(x_3, 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); +lean_inc(x_47); +x_49 = lean_apply_5(x_4, x_5, x_47, x_45, x_8, x_48); +if (lean_obj_tag(x_49) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_18, 1); -lean_inc(x_35); -lean_dec(x_18); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_MessageData_toString(x_36, x_35); -if (lean_obj_tag(x_37) == 0) +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_50 = lean_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_st_ref_get(x_47, x_51); +lean_dec(x_47); +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_52); + x_55 = lean_box(0); +} +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_50); +lean_ctor_set(x_56, 1, x_53); +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 { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +lean_dec(x_47); +x_58 = lean_ctor_get(x_49, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_49, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_60 = x_49; +} else { + lean_dec_ref(x_49); + x_60 = lean_box(0); +} +if (lean_is_scalar(x_60)) { + x_61 = lean_alloc_ctor(1, 2, 0); +} else { + x_61 = x_60; +} +lean_ctor_set(x_61, 0, x_58); +lean_ctor_set(x_61, 1, x_59); +return x_61; +} +} +} +} +static lean_object* _init_l_Lean_Meta_MetaM_toIO___rarg___closed__1() { +_start: { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set_tag(x_37, 1); -lean_ctor_set(x_37, 0, x_40); -return x_37; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("internal exception #", 20); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_MetaM_toIO___rarg___closed__2() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_37, 0); -x_42 = lean_ctor_get(x_37, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_37); -x_43 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_43, 0, x_41); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -return x_44; +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -uint8_t x_45; -x_45 = !lean_is_exclusive(x_37); -if (x_45 == 0) +lean_object* x_7; lean_object* x_40; lean_object* x_41; lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_62 = lean_io_get_num_heartbeats(x_6); +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_is_exclusive(x_2); +if (x_65 == 0) { -return x_37; +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; uint8_t x_78; +x_66 = lean_ctor_get(x_2, 2); +x_67 = lean_ctor_get(x_2, 8); +lean_dec(x_67); +lean_inc(x_66); +lean_ctor_set(x_2, 8, x_63); +x_68 = lean_st_mk_ref(x_3, x_64); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_Meta_MetaM_toIO___rarg___closed__2; +x_72 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_66, x_71); +x_73 = lean_st_ref_get(x_69, x_70); +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, 0); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Kernel_isDiagnosticsEnabled(x_76); +lean_dec(x_76); +if (x_77 == 0) +{ +if (x_72 == 0) +{ +uint8_t x_141; +x_141 = 1; +x_78 = x_141; +goto block_140; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_37, 0); -x_47 = lean_ctor_get(x_37, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_37); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; -} +uint8_t x_142; +x_142 = 0; +x_78 = x_142; +goto block_140; } } else { -uint8_t x_49; -x_49 = !lean_is_exclusive(x_18); -if (x_49 == 0) +if (x_72 == 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; -x_50 = lean_ctor_get(x_18, 0); -lean_dec(x_50); -x_51 = lean_ctor_get(x_34, 0); -lean_inc(x_51); -lean_dec(x_34); -x_52 = l___private_Init_Data_Repr_0__Nat_reprFast(x_51); -x_53 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; -x_54 = lean_string_append(x_53, x_52); -lean_dec(x_52); -x_55 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_18, 0, x_55); -return x_18; +uint8_t x_143; +x_143 = 0; +x_78 = x_143; +goto block_140; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_56 = lean_ctor_get(x_18, 1); -lean_inc(x_56); -lean_dec(x_18); -x_57 = lean_ctor_get(x_34, 0); -lean_inc(x_57); -lean_dec(x_34); -x_58 = l___private_Init_Data_Repr_0__Nat_reprFast(x_57); -x_59 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; -x_60 = lean_string_append(x_59, x_58); -lean_dec(x_58); -x_61 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_61, 0, x_60); -x_62 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_56); -return x_62; -} +uint8_t x_144; +x_144 = 1; +x_78 = x_144; +goto block_140; } } -} -else +block_140: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; uint8_t x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_63 = lean_ctor_get(x_2, 0); -x_64 = lean_ctor_get(x_2, 1); -x_65 = lean_ctor_get(x_2, 2); -x_66 = lean_ctor_get(x_2, 3); -x_67 = lean_ctor_get(x_2, 4); -x_68 = lean_ctor_get(x_2, 5); -x_69 = lean_ctor_get(x_2, 6); -x_70 = lean_ctor_get(x_2, 7); -x_71 = lean_ctor_get(x_2, 9); -x_72 = lean_ctor_get(x_2, 10); -x_73 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); -x_74 = lean_ctor_get_uint8(x_2, sizeof(void*)*11 + 1); -lean_inc(x_72); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_inc(x_67); -lean_inc(x_66); -lean_inc(x_65); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_2); -x_75 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_75, 0, x_63); -lean_ctor_set(x_75, 1, x_64); -lean_ctor_set(x_75, 2, x_65); -lean_ctor_set(x_75, 3, x_66); -lean_ctor_set(x_75, 4, x_67); -lean_ctor_set(x_75, 5, x_68); -lean_ctor_set(x_75, 6, x_69); -lean_ctor_set(x_75, 7, x_70); -lean_ctor_set(x_75, 8, x_8); -lean_ctor_set(x_75, 9, x_71); -lean_ctor_set(x_75, 10, x_72); -lean_ctor_set_uint8(x_75, sizeof(void*)*11, x_73); -lean_ctor_set_uint8(x_75, sizeof(void*)*11 + 1, x_74); -x_76 = lean_st_mk_ref(x_3, x_9); -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_mk_ref(x_5, x_78); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_79 = lean_st_ref_take(x_69, x_75); 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); -lean_inc(x_77); -lean_inc(x_80); -x_82 = lean_apply_5(x_1, x_4, x_80, x_75, x_77, x_81); -if (lean_obj_tag(x_82) == 0) +x_82 = !lean_is_exclusive(x_80); +if (x_82 == 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; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); -lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_st_ref_get(x_80, x_84); -lean_dec(x_80); -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_st_ref_get(x_77, x_87); -lean_dec(x_77); -x_89 = lean_ctor_get(x_88, 0); -lean_inc(x_89); -x_90 = lean_ctor_get(x_88, 1); -lean_inc(x_90); -if (lean_is_exclusive(x_88)) { - lean_ctor_release(x_88, 0); - lean_ctor_release(x_88, 1); - x_91 = x_88; -} else { - lean_dec_ref(x_88); - x_91 = lean_box(0); -} -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_89); -lean_ctor_set(x_92, 1, x_86); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_83); -lean_ctor_set(x_93, 1, x_92); -if (lean_is_scalar(x_91)) { - x_94 = lean_alloc_ctor(0, 2, 0); -} else { - x_94 = x_91; -} -lean_ctor_set(x_94, 0, x_93); -lean_ctor_set(x_94, 1, x_90); -return x_94; +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_80, 0); +x_84 = lean_ctor_get(x_80, 4); +lean_dec(x_84); +x_85 = l_Lean_Kernel_enableDiag(x_83, x_72); +x_86 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; +lean_ctor_set(x_80, 4, x_86); +lean_ctor_set(x_80, 0, x_85); +x_87 = lean_st_ref_set(x_69, x_80, x_81); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_box(0); +lean_inc(x_69); +x_90 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_66, x_72, x_5, x_1, x_4, x_89, x_2, x_69, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_st_ref_get(x_69, x_92); +lean_dec(x_69); +x_94 = !lean_is_exclusive(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_93, 0); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_91); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_93, 0, x_96); +x_7 = x_93; +goto block_39; } else { -lean_object* x_95; -lean_dec(x_80); -lean_dec(x_77); -x_95 = lean_ctor_get(x_82, 0); -lean_inc(x_95); -if (lean_obj_tag(x_95) == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_82, 1); -lean_inc(x_96); -lean_dec(x_82); -x_97 = lean_ctor_get(x_95, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_93, 0); +x_98 = lean_ctor_get(x_93, 1); +lean_inc(x_98); lean_inc(x_97); -lean_dec(x_95); -x_98 = l_Lean_MessageData_toString(x_97, x_96); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_101 = x_98; -} else { - lean_dec_ref(x_98); - x_101 = lean_box(0); -} -x_102 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_102, 0, x_99); -if (lean_is_scalar(x_101)) { - x_103 = lean_alloc_ctor(1, 2, 0); -} else { - x_103 = x_101; - lean_ctor_set_tag(x_103, 1); +lean_dec(x_93); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_91); +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_98); +x_7 = x_100; +goto block_39; } -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_100); -return x_103; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_106 = x_98; -} else { - lean_dec_ref(x_98); - x_106 = lean_box(0); -} -if (lean_is_scalar(x_106)) { - x_107 = lean_alloc_ctor(1, 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; +lean_object* x_101; lean_object* x_102; +lean_dec(x_69); +x_101 = lean_ctor_get(x_90, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_90, 1); +lean_inc(x_102); +lean_dec(x_90); +x_40 = x_101; +x_41 = x_102; +goto block_61; } } else { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_108 = lean_ctor_get(x_82, 1); +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_103 = lean_ctor_get(x_80, 0); +x_104 = lean_ctor_get(x_80, 1); +x_105 = lean_ctor_get(x_80, 2); +x_106 = lean_ctor_get(x_80, 3); +x_107 = lean_ctor_get(x_80, 5); +x_108 = lean_ctor_get(x_80, 6); lean_inc(x_108); -if (lean_is_exclusive(x_82)) { - lean_ctor_release(x_82, 0); - lean_ctor_release(x_82, 1); - x_109 = x_82; +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_80); +x_109 = l_Lean_Kernel_enableDiag(x_103, x_72); +x_110 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; +x_111 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_104); +lean_ctor_set(x_111, 2, x_105); +lean_ctor_set(x_111, 3, x_106); +lean_ctor_set(x_111, 4, x_110); +lean_ctor_set(x_111, 5, x_107); +lean_ctor_set(x_111, 6, x_108); +x_112 = lean_st_ref_set(x_69, x_111, x_81); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +lean_dec(x_112); +x_114 = lean_box(0); +lean_inc(x_69); +x_115 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_66, x_72, x_5, x_1, x_4, x_114, x_2, x_69, x_113); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_st_ref_get(x_69, x_117); +lean_dec(x_69); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_121 = x_118; } else { - lean_dec_ref(x_82); - x_109 = lean_box(0); + lean_dec_ref(x_118); + x_121 = lean_box(0); } -x_110 = lean_ctor_get(x_95, 0); -lean_inc(x_110); -lean_dec(x_95); -x_111 = l___private_Init_Data_Repr_0__Nat_reprFast(x_110); -x_112 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; -x_113 = lean_string_append(x_112, x_111); -lean_dec(x_111); -x_114 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_114, 0, x_113); -if (lean_is_scalar(x_109)) { - x_115 = lean_alloc_ctor(1, 2, 0); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_116); +lean_ctor_set(x_122, 1, x_119); +if (lean_is_scalar(x_121)) { + x_123 = lean_alloc_ctor(0, 2, 0); } else { - x_115 = x_109; -} -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_108); -return x_115; + x_123 = x_121; } +lean_ctor_set(x_123, 0, x_122); +lean_ctor_set(x_123, 1, x_120); +x_7 = x_123; +goto block_39; } +else +{ +lean_object* x_124; lean_object* x_125; +lean_dec(x_69); +x_124 = lean_ctor_get(x_115, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_115, 1); +lean_inc(x_125); +lean_dec(x_115); +x_40 = x_124; +x_41 = x_125; +goto block_61; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_MetaM_toIO___rarg), 6, 0); -return x_2; -} +lean_object* x_126; lean_object* x_127; +x_126 = lean_box(0); +lean_inc(x_69); +x_127 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_66, x_72, x_5, x_1, x_4, x_126, x_2, x_69, x_75); +if (lean_obj_tag(x_127) == 0) +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = lean_st_ref_get(x_69, x_129); +lean_dec(x_69); +x_131 = !lean_is_exclusive(x_130); +if (x_131 == 0) +{ +lean_object* x_132; lean_object* x_133; +x_132 = lean_ctor_get(x_130, 0); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_128); +lean_ctor_set(x_133, 1, x_132); +lean_ctor_set(x_130, 0, x_133); +x_7 = x_130; +goto block_39; } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__1() { -_start: +else { -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; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Context_config___default___closed__1; -x_3 = l_Lean_Meta_Context_lctx___default___closed__2; -x_4 = l_Lean_Meta_ParamInfo_backDeps___default___closed__1; -x_5 = lean_unsigned_to_nat(0u); -x_6 = 0; -x_7 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_7, 0, x_2); -lean_ctor_set(x_7, 1, x_3); -lean_ctor_set(x_7, 2, x_4); -lean_ctor_set(x_7, 3, x_1); -lean_ctor_set(x_7, 4, x_5); -lean_ctor_set(x_7, 5, x_1); -lean_ctor_set_uint8(x_7, sizeof(void*)*6, x_6); -lean_ctor_set_uint8(x_7, sizeof(void*)*6 + 1, x_6); -return x_7; +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_134 = lean_ctor_get(x_130, 0); +x_135 = lean_ctor_get(x_130, 1); +lean_inc(x_135); +lean_inc(x_134); +lean_dec(x_130); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_128); +lean_ctor_set(x_136, 1, x_134); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_135); +x_7 = x_137; +goto block_39; } } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2() { -_start: +else { -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_State_mctx___default___closed__3; -x_3 = l_Lean_Meta_instInhabitedCache___closed__4; -x_4 = l_Lean_Meta_State_postponed___default___closed__3; -x_5 = l_Lean_Meta_instInhabitedDiagnostics___closed__1; -x_6 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_6, 0, x_2); -lean_ctor_set(x_6, 1, x_3); -lean_ctor_set(x_6, 2, x_1); -lean_ctor_set(x_6, 3, x_4); -lean_ctor_set(x_6, 4, x_5); -return x_6; +lean_object* x_138; lean_object* x_139; +lean_dec(x_69); +x_138 = lean_ctor_get(x_127, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_127, 1); +lean_inc(x_139); +lean_dec(x_127); +x_40 = x_138; +x_41 = x_139; +goto block_61; } } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 0); -return x_1; } } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -x_2 = l_Lean_Meta_ParamInfo_backDeps___default___closed__1; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; -} +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; uint8_t x_155; uint8_t x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; uint8_t x_167; uint8_t x_168; +x_145 = lean_ctor_get(x_2, 0); +x_146 = lean_ctor_get(x_2, 1); +x_147 = lean_ctor_get(x_2, 2); +x_148 = lean_ctor_get(x_2, 3); +x_149 = lean_ctor_get(x_2, 4); +x_150 = lean_ctor_get(x_2, 5); +x_151 = lean_ctor_get(x_2, 6); +x_152 = lean_ctor_get(x_2, 7); +x_153 = lean_ctor_get(x_2, 9); +x_154 = lean_ctor_get(x_2, 10); +x_155 = lean_ctor_get_uint8(x_2, sizeof(void*)*11); +x_156 = lean_ctor_get_uint8(x_2, sizeof(void*)*11 + 1); +lean_inc(x_154); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_inc(x_146); +lean_inc(x_145); +lean_dec(x_2); +lean_inc(x_147); +x_157 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_157, 0, x_145); +lean_ctor_set(x_157, 1, x_146); +lean_ctor_set(x_157, 2, x_147); +lean_ctor_set(x_157, 3, x_148); +lean_ctor_set(x_157, 4, x_149); +lean_ctor_set(x_157, 5, x_150); +lean_ctor_set(x_157, 6, x_151); +lean_ctor_set(x_157, 7, x_152); +lean_ctor_set(x_157, 8, x_63); +lean_ctor_set(x_157, 9, x_153); +lean_ctor_set(x_157, 10, x_154); +lean_ctor_set_uint8(x_157, sizeof(void*)*11, x_155); +lean_ctor_set_uint8(x_157, sizeof(void*)*11 + 1, x_156); +x_158 = lean_st_mk_ref(x_3, x_64); +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_Lean_Meta_MetaM_toIO___rarg___closed__2; +x_162 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_147, x_161); +x_163 = lean_st_ref_get(x_159, x_160); +x_164 = lean_ctor_get(x_163, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_163, 1); +lean_inc(x_165); +lean_dec(x_163); +x_166 = lean_ctor_get(x_164, 0); +lean_inc(x_166); +lean_dec(x_164); +x_167 = l_Lean_Kernel_isDiagnosticsEnabled(x_166); +lean_dec(x_166); +if (x_167 == 0) +{ +if (x_162 == 0) +{ +uint8_t x_209; +x_209 = 1; +x_168 = x_209; +goto block_208; } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_maxRecDepth; -return x_1; +uint8_t x_210; +x_210 = 0; +x_168 = x_210; +goto block_208; } } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__6() { -_start: +else { -lean_object* x_1; -x_1 = l_Lean_diag; -return x_1; -} +if (x_162 == 0) +{ +uint8_t x_211; +x_211 = 0; +x_168 = x_211; +goto block_208; } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_firstFrontendMacroScope; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_nat_add(x_1, x_2); -return x_3; +uint8_t x_212; +x_212 = 1; +x_168 = x_212; +goto block_208; } } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__8() { -_start: +block_208: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_uniq", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9() { -_start: +if (x_168 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_169 = lean_st_ref_take(x_159, x_165); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +lean_dec(x_169); +x_172 = lean_ctor_get(x_170, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_170, 1); +lean_inc(x_173); +x_174 = lean_ctor_get(x_170, 2); +lean_inc(x_174); +x_175 = lean_ctor_get(x_170, 3); +lean_inc(x_175); +x_176 = lean_ctor_get(x_170, 5); +lean_inc(x_176); +x_177 = lean_ctor_get(x_170, 6); +lean_inc(x_177); +if (lean_is_exclusive(x_170)) { + lean_ctor_release(x_170, 0); + lean_ctor_release(x_170, 1); + lean_ctor_release(x_170, 2); + lean_ctor_release(x_170, 3); + lean_ctor_release(x_170, 4); + lean_ctor_release(x_170, 5); + lean_ctor_release(x_170, 6); + x_178 = x_170; +} else { + lean_dec_ref(x_170); + x_178 = lean_box(0); } +x_179 = l_Lean_Kernel_enableDiag(x_172, x_162); +x_180 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; +if (lean_is_scalar(x_178)) { + x_181 = lean_alloc_ctor(0, 7, 0); +} else { + x_181 = x_178; } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__10() { -_start: +lean_ctor_set(x_181, 0, x_179); +lean_ctor_set(x_181, 1, x_173); +lean_ctor_set(x_181, 2, x_174); +lean_ctor_set(x_181, 3, x_175); +lean_ctor_set(x_181, 4, x_180); +lean_ctor_set(x_181, 5, x_176); +lean_ctor_set(x_181, 6, x_177); +x_182 = lean_st_ref_set(x_159, x_181, x_171); +x_183 = lean_ctor_get(x_182, 1); +lean_inc(x_183); +lean_dec(x_182); +x_184 = lean_box(0); +lean_inc(x_159); +x_185 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_147, x_162, x_5, x_1, x_4, x_184, x_157, x_159, x_183); +if (lean_obj_tag(x_185) == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9; -x_2 = lean_unsigned_to_nat(1u); -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_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_188 = lean_st_ref_get(x_159, x_187); +lean_dec(x_159); +x_189 = lean_ctor_get(x_188, 0); +lean_inc(x_189); +x_190 = lean_ctor_get(x_188, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_188)) { + lean_ctor_release(x_188, 0); + lean_ctor_release(x_188, 1); + x_191 = x_188; +} else { + lean_dec_ref(x_188); + x_191 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__11() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 1; -x_2 = l_Lean_Meta_State_mctx___default___closed__2; -x_3 = l_Lean_Meta_State_postponed___default___closed__3; -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; +x_192 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_192, 0, x_186); +lean_ctor_set(x_192, 1, x_189); +if (lean_is_scalar(x_191)) { + x_193 = lean_alloc_ctor(0, 2, 0); +} else { + x_193 = x_191; } +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_190); +x_7 = x_193; +goto block_39; } -static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__12() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 7); -return x_1; +lean_object* x_194; lean_object* x_195; +lean_dec(x_159); +x_194 = lean_ctor_get(x_185, 0); +lean_inc(x_194); +x_195 = lean_ctor_get(x_185, 1); +lean_inc(x_195); +lean_dec(x_185); +x_40 = x_194; +x_41 = x_195; +goto block_61; } } -LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___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) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; 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_29 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5; -x_30 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_3, x_29); -x_31 = lean_box(0); -x_32 = l_Lean_Core_getMaxHeartbeats(x_3); -x_33 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__6; -x_34 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_3, x_33); -x_35 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__7; -x_36 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__10; -x_37 = l_Lean_Meta_State_postponed___default___closed__3; -x_38 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; -x_39 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__11; -x_40 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_40, 0, x_2); -lean_ctor_set(x_40, 1, x_35); -lean_ctor_set(x_40, 2, x_36); -lean_ctor_set(x_40, 3, x_37); -lean_ctor_set(x_40, 4, x_38); -lean_ctor_set(x_40, 5, x_37); -lean_ctor_set(x_40, 6, x_39); -x_41 = lean_io_get_num_heartbeats(x_6); -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_Meta_instMetaEvalMetaM___rarg___closed__12; -x_45 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__4; -x_46 = lean_unsigned_to_nat(0u); -x_47 = lean_box(0); -x_48 = lean_box(0); -x_49 = l_Lean_firstFrontendMacroScope; -x_50 = 0; -lean_inc(x_3); -x_51 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_51, 0, x_44); -lean_ctor_set(x_51, 1, x_45); -lean_ctor_set(x_51, 2, x_3); -lean_ctor_set(x_51, 3, x_46); -lean_ctor_set(x_51, 4, x_30); -lean_ctor_set(x_51, 5, x_47); -lean_ctor_set(x_51, 6, x_48); -lean_ctor_set(x_51, 7, x_31); -lean_ctor_set(x_51, 8, x_42); -lean_ctor_set(x_51, 9, x_32); -lean_ctor_set(x_51, 10, x_49); -lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_50); -lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_34); -x_52 = lean_st_mk_ref(x_40, x_43); -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_Meta_instMetaEvalMetaM___rarg___closed__2; -x_56 = lean_st_mk_ref(x_55, x_54); -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_Meta_instMetaEvalMetaM___rarg___closed__1; -lean_inc(x_53); -lean_inc(x_51); -lean_inc(x_57); -x_60 = lean_apply_5(x_4, x_59, x_57, x_51, x_53, x_58); -if (lean_obj_tag(x_60) == 0) +else { -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_st_ref_get(x_57, x_62); -lean_dec(x_57); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -lean_dec(x_63); -lean_inc(x_53); -x_65 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_51, x_53, x_64); -if (lean_obj_tag(x_65) == 0) +lean_object* x_196; lean_object* x_197; +x_196 = lean_box(0); +lean_inc(x_159); +x_197 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_147, x_162, x_5, x_1, x_4, x_196, x_157, x_159, x_165); +if (lean_obj_tag(x_197) == 0) { -lean_object* 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; -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = lean_st_ref_get(x_53, x_66); -lean_dec(x_53); -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, 0); -lean_inc(x_70); -lean_dec(x_68); -x_71 = 1; -x_72 = lean_box(x_71); -x_73 = lean_apply_5(x_1, x_70, x_3, x_61, x_72, x_69); -return x_73; +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +x_198 = lean_ctor_get(x_197, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_197, 1); +lean_inc(x_199); +lean_dec(x_197); +x_200 = lean_st_ref_get(x_159, x_199); +lean_dec(x_159); +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +if (lean_is_exclusive(x_200)) { + lean_ctor_release(x_200, 0); + lean_ctor_release(x_200, 1); + x_203 = x_200; +} else { + lean_dec_ref(x_200); + x_203 = lean_box(0); } -else -{ -lean_object* x_74; lean_object* x_75; -lean_dec(x_61); -lean_dec(x_53); -lean_dec(x_3); -lean_dec(x_1); -x_74 = lean_ctor_get(x_65, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_65, 1); -lean_inc(x_75); -lean_dec(x_65); -x_7 = x_74; -x_8 = x_75; -goto block_28; +x_204 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_204, 0, x_198); +lean_ctor_set(x_204, 1, x_201); +if (lean_is_scalar(x_203)) { + x_205 = lean_alloc_ctor(0, 2, 0); +} else { + x_205 = x_203; } +lean_ctor_set(x_205, 0, x_204); +lean_ctor_set(x_205, 1, x_202); +x_7 = x_205; +goto block_39; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -lean_dec(x_57); -lean_dec(x_3); -lean_dec(x_1); -x_76 = lean_ctor_get(x_60, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_60, 1); -lean_inc(x_77); -lean_dec(x_60); -x_78 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_51, x_53, x_77); -if (lean_obj_tag(x_78) == 0) -{ -lean_object* x_79; -x_79 = lean_ctor_get(x_78, 1); -lean_inc(x_79); -lean_dec(x_78); -x_7 = x_76; -x_8 = x_79; -goto block_28; +lean_object* x_206; lean_object* x_207; +lean_dec(x_159); +x_206 = lean_ctor_get(x_197, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_197, 1); +lean_inc(x_207); +lean_dec(x_197); +x_40 = x_206; +x_41 = x_207; +goto block_61; } -else -{ -lean_object* x_80; lean_object* x_81; -lean_dec(x_76); -x_80 = lean_ctor_get(x_78, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_78, 1); -lean_inc(x_81); -lean_dec(x_78); -x_7 = x_80; -x_8 = x_81; -goto block_28; } } -block_28: +} +block_39: { if (lean_obj_tag(x_7) == 0) { -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 1); +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 0); lean_inc(x_9); -lean_dec(x_7); -x_10 = l_Lean_MessageData_toString(x_9, x_8); -if (lean_obj_tag(x_10) == 0) +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = !lean_is_exclusive(x_8); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set_tag(x_10, 1); -lean_ctor_set(x_10, 0, x_13); -return x_10; -} -else +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_8, 1); +x_14 = lean_ctor_get(x_8, 0); +lean_dec(x_14); +x_15 = !lean_is_exclusive(x_9); +if (x_15 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_10, 0); -x_15 = lean_ctor_get(x_10, 1); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_10); -x_16 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_16, 0, 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_16; +x_16 = lean_ctor_get(x_9, 0); +lean_ctor_set(x_9, 0, x_13); +lean_ctor_set(x_8, 1, x_9); +lean_ctor_set(x_8, 0, x_16); +return x_7; } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_10); -if (x_18 == 0) -{ -return x_10; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_13); +lean_ctor_set(x_19, 1, x_18); +lean_ctor_set(x_8, 1, x_19); +lean_ctor_set(x_8, 0, x_17); +return x_7; +} } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_10, 0); -x_20 = lean_ctor_get(x_10, 1); +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_8, 1); lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_10); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +lean_dec(x_8); +x_21 = lean_ctor_get(x_9, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_9, 1); +lean_inc(x_22); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_23 = x_9; +} else { + lean_dec_ref(x_9); + x_23 = lean_box(0); +} +if (lean_is_scalar(x_23)) { + x_24 = lean_alloc_ctor(0, 2, 0); +} else { + x_24 = x_23; } +lean_ctor_set(x_24, 0, x_20); +lean_ctor_set(x_24, 1, x_22); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_21); +lean_ctor_set(x_25, 1, x_24); +lean_ctor_set(x_7, 0, x_25); +return x_7; } } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_22 = lean_ctor_get(x_7, 0); -lean_inc(x_22); +lean_object* x_26; lean_object* x_27; lean_object* x_28; 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_ctor_get(x_7, 1); +lean_inc(x_26); lean_dec(x_7); -x_23 = l___private_Init_Data_Repr_0__Nat_reprFast(x_22); -x_24 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; -x_25 = lean_string_append(x_24, x_23); -lean_dec(x_23); -x_26 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_26, 0, x_25); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_8); -return x_27; +x_27 = lean_ctor_get(x_8, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_8)) { + lean_ctor_release(x_8, 0); + lean_ctor_release(x_8, 1); + x_28 = x_8; +} else { + lean_dec_ref(x_8); + x_28 = lean_box(0); } +x_29 = lean_ctor_get(x_9, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_9, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_31 = x_9; +} else { + lean_dec_ref(x_9); + x_31 = lean_box(0); } +if (lean_is_scalar(x_31)) { + x_32 = lean_alloc_ctor(0, 2, 0); +} else { + x_32 = x_31; } +lean_ctor_set(x_32, 0, x_27); +lean_ctor_set(x_32, 1, x_30); +if (lean_is_scalar(x_28)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_28; } -LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_instMetaEvalMetaM___rarg___boxed), 6, 0); -return x_2; +lean_ctor_set(x_33, 0, x_29); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_26); +return x_34; } } -LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +else { -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_5); -lean_dec(x_5); -x_8 = l_Lean_Meta_instMetaEvalMetaM___rarg(x_1, x_2, x_3, x_4, x_7, x_6); -return x_8; -} -} -static lean_object* _init_l_Lean_Meta_throwIsDefEqStuck___rarg___closed__1() { -_start: +uint8_t x_35; +x_35 = !lean_is_exclusive(x_7); +if (x_35 == 0) { -lean_object* x_1; -x_1 = l_Lean_Meta_isDefEqStuckExceptionId; -return x_1; -} +return x_7; } -static lean_object* _init_l_Lean_Meta_throwIsDefEqStuck___rarg___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_throwIsDefEqStuck___rarg___closed__1; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_7, 0); +x_37 = lean_ctor_get(x_7, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_7); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } -LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg(lean_object* x_1) { -_start: -{ -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_throwIsDefEqStuck___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; } +block_61: +{ +if (lean_obj_tag(x_40) == 0) +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); +x_43 = l_Lean_MessageData_toString(x_42, x_41); +if (lean_obj_tag(x_43) == 0) +{ +uint8_t x_44; +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_43, 0); +x_46 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_46, 0, x_45); +lean_ctor_set_tag(x_43, 1); +lean_ctor_set(x_43, 0, x_46); +x_7 = x_43; +goto block_39; } -LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck(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; -x_6 = lean_alloc_closure((void*)(l_Lean_Meta_throwIsDefEqStuck___rarg), 1, 0); -return x_6; +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 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_49, 0, x_47); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_48); +x_7 = x_50; +goto block_39; } } -LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___boxed(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; -x_6 = l_Lean_Meta_throwIsDefEqStuck(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_6; -} +uint8_t x_51; +x_51 = !lean_is_exclusive(x_43); +if (x_51 == 0) +{ +x_7 = x_43; +goto block_39; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Meta", 4); -return x_1; +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); +x_7 = x_54; +goto block_39; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__2() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_55 = lean_ctor_get(x_40, 0); +lean_inc(x_55); +lean_dec(x_40); +x_56 = l___private_Init_Data_Repr_0__Nat_reprFast(x_55); +x_57 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; +x_58 = lean_string_append(x_57, x_56); +lean_dec(x_56); +x_59 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_59, 0, x_58); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_41); +x_7 = x_60; +goto block_39; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean", 4); -return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__4() { +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_MetaM_toIO___rarg), 6, 0); +return x_2; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__5() { +LEAN_EXPORT lean_object* l_Lean_Meta_MetaM_toIO___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) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_11; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__6() { +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___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_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("initFn", 6); -return x_1; -} +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; +x_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +lean_ctor_set(x_8, 4, x_2); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_3); +x_14 = lean_st_mk_ref(x_4, 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); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_15); +x_17 = lean_apply_5(x_5, x_6, x_15, x_8, x_9, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_st_ref_get(x_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 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_8, x_9, x_21); +if (lean_obj_tag(x_22) == 0) +{ +uint8_t x_23; +x_23 = !lean_is_exclusive(x_22); +if (x_23 == 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; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__7() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__6; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_22, 1); +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; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__8() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_@", 2); -return x_1; -} +uint8_t x_27; +lean_dec(x_18); +x_27 = !lean_is_exclusive(x_22); +if (x_27 == 0) +{ +return x_22; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__9() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__7; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__8; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_22, 0); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_22); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__10() { -_start: +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__9; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__3; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; -} +lean_object* x_31; lean_object* x_32; lean_object* x_33; +lean_dec(x_15); +x_31 = lean_ctor_get(x_17, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_17, 1); +lean_inc(x_32); +lean_dec(x_17); +x_33 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_8, x_9, x_32); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +lean_ctor_set_tag(x_33, 1); +lean_ctor_set(x_33, 0, x_31); +return x_33; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__11() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +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(1, 2, 0); +lean_ctor_set(x_37, 0, x_31); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__12() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Basic", 5); -return x_1; -} +uint8_t x_38; +lean_dec(x_31); +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) +{ +return x_33; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__13() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__11; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__12; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_33); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__14() { -_start: +} +} +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_hyg", 4); -return x_1; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_42 = lean_ctor_get(x_8, 0); +x_43 = lean_ctor_get(x_8, 1); +x_44 = lean_ctor_get(x_8, 3); +x_45 = lean_ctor_get(x_8, 5); +x_46 = lean_ctor_get(x_8, 6); +x_47 = lean_ctor_get(x_8, 7); +x_48 = lean_ctor_get(x_8, 8); +x_49 = lean_ctor_get(x_8, 9); +x_50 = lean_ctor_get(x_8, 10); +x_51 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_8); +x_52 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_52, 0, x_42); +lean_ctor_set(x_52, 1, x_43); +lean_ctor_set(x_52, 2, x_1); +lean_ctor_set(x_52, 3, x_44); +lean_ctor_set(x_52, 4, x_2); +lean_ctor_set(x_52, 5, x_45); +lean_ctor_set(x_52, 6, x_46); +lean_ctor_set(x_52, 7, x_47); +lean_ctor_set(x_52, 8, x_48); +lean_ctor_set(x_52, 9, x_49); +lean_ctor_set(x_52, 10, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*11, x_51); +lean_ctor_set_uint8(x_52, sizeof(void*)*11 + 1, x_3); +x_53 = lean_st_mk_ref(x_4, x_10); +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); +lean_inc(x_9); +lean_inc(x_52); +lean_inc(x_54); +x_56 = lean_apply_5(x_5, x_6, x_54, x_52, x_9, x_55); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_56, 1); +lean_inc(x_58); +lean_dec(x_56); +x_59 = lean_st_ref_get(x_54, x_58); +lean_dec(x_54); +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); +x_61 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_52, x_9, 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_61, 1); +lean_inc(x_62); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_63 = x_61; +} else { + lean_dec_ref(x_61); + x_63 = lean_box(0); } +if (lean_is_scalar(x_63)) { + x_64 = lean_alloc_ctor(0, 2, 0); +} else { + x_64 = x_63; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__15() { -_start: +lean_ctor_set(x_64, 0, x_57); +lean_ctor_set(x_64, 1, x_62); +return x_64; +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__13; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__14; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_57); +x_65 = lean_ctor_get(x_61, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_61, 1); +lean_inc(x_66); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_67 = x_61; +} else { + lean_dec_ref(x_61); + x_67 = lean_box(0); } +if (lean_is_scalar(x_67)) { + x_68 = lean_alloc_ctor(1, 2, 0); +} else { + x_68 = x_67; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__16() { -_start: +lean_ctor_set(x_68, 0, x_65); +lean_ctor_set(x_68, 1, x_66); +return x_68; +} +} +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__15; -x_2 = lean_unsigned_to_nat(1758u); -x_3 = l_Lean_Name_num___override(x_1, x_2); -return x_3; +lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_54); +x_69 = lean_ctor_get(x_56, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_56, 1); +lean_inc(x_70); +lean_dec(x_56); +x_71 = l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___spec__1(x_52, x_9, x_70); +if (lean_obj_tag(x_71) == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_71, 1); +lean_inc(x_72); +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); +} +if (lean_is_scalar(x_73)) { + x_74 = lean_alloc_ctor(1, 2, 0); +} else { + x_74 = x_73; + lean_ctor_set_tag(x_74, 1); } +lean_ctor_set(x_74, 0, x_69); +lean_ctor_set(x_74, 1, x_72); +return x_74; } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__17() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("debug", 5); -return x_1; +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +lean_dec(x_69); +x_75 = lean_ctor_get(x_71, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_71, 1); +lean_inc(x_76); +if (lean_is_exclusive(x_71)) { + lean_ctor_release(x_71, 0); + lean_ctor_release(x_71, 1); + x_77 = x_71; +} else { + lean_dec_ref(x_71); + x_77 = lean_box(0); +} +if (lean_is_scalar(x_77)) { + x_78 = lean_alloc_ctor(1, 2, 0); +} else { + x_78 = x_77; +} +lean_ctor_set(x_78, 0, x_75); +lean_ctor_set(x_78, 1, x_76); +return x_78; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__18() { +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__17; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +uint8_t x_11; +lean_dec(x_7); +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; uint8_t x_20; uint8_t x_21; +x_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_inc(x_15); +lean_inc(x_1); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); +x_16 = lean_st_ref_get(x_9, x_10); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Kernel_isDiagnosticsEnabled(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +if (x_2 == 0) +{ +uint8_t x_48; +x_48 = 1; +x_21 = x_48; +goto block_47; } +else +{ +uint8_t x_49; +x_49 = 0; +x_21 = x_49; +goto block_47; } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758_(lean_object* x_1) { -_start: +} +else { -lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__2; -x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__16; -x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); -if (lean_obj_tag(x_5) == 0) +if (x_2 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__18; -x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); -return x_8; +uint8_t x_50; +x_50 = 0; +x_21 = x_50; +goto block_47; } else { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_5); -if (x_9 == 0) +uint8_t x_51; +x_51 = 1; +x_21 = x_51; +goto block_47; +} +} +block_47: { -return x_5; +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_st_ref_take(x_9, x_18); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; 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_23, 0); +x_27 = lean_ctor_get(x_23, 4); +lean_dec(x_27); +x_28 = l_Lean_Kernel_enableDiag(x_26, x_2); +lean_ctor_set(x_23, 4, x_6); +lean_ctor_set(x_23, 0, x_28); +x_29 = lean_st_ref_set(x_9, x_23, x_24); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_box(0); +x_32 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_15, x_2, x_3, x_4, x_5, x_31, x_8, x_9, x_30); +return x_32; } else { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_5, 0); -x_11 = lean_ctor_get(x_5, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_5); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_33 = lean_ctor_get(x_23, 0); +x_34 = lean_ctor_get(x_23, 1); +x_35 = lean_ctor_get(x_23, 2); +x_36 = lean_ctor_get(x_23, 3); +x_37 = lean_ctor_get(x_23, 5); +x_38 = lean_ctor_get(x_23, 6); +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_23); +x_39 = l_Lean_Kernel_enableDiag(x_33, x_2); +x_40 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_34); +lean_ctor_set(x_40, 2, x_35); +lean_ctor_set(x_40, 3, x_36); +lean_ctor_set(x_40, 4, x_6); +lean_ctor_set(x_40, 5, x_37); +lean_ctor_set(x_40, 6, x_38); +x_41 = lean_st_ref_set(x_9, x_40, x_24); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_box(0); +x_44 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_15, x_2, x_3, x_4, x_5, x_43, x_8, x_9, x_42); +return x_44; } } +else +{ +lean_object* x_45; lean_object* x_46; +lean_dec(x_6); +x_45 = lean_box(0); +x_46 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_15, x_2, x_3, x_4, x_5, x_45, x_8, x_9, x_18); +return x_46; +} } } -LEAN_EXPORT lean_object* l_Lean_Meta_liftMetaM___rarg(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_apply_2(x_1, lean_box(0), x_2); -return x_3; +lean_object* x_52; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; uint8_t x_70; +x_52 = lean_ctor_get(x_8, 0); +x_53 = lean_ctor_get(x_8, 1); +x_54 = lean_ctor_get(x_8, 3); +x_55 = lean_ctor_get(x_8, 5); +x_56 = lean_ctor_get(x_8, 6); +x_57 = lean_ctor_get(x_8, 7); +x_58 = lean_ctor_get(x_8, 8); +x_59 = lean_ctor_get(x_8, 9); +x_60 = lean_ctor_get(x_8, 10); +x_61 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +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_inc(x_52); +lean_dec(x_8); +x_62 = l_Lean_Meta_MetaM_toIO___rarg___lambda__1___closed__1; +x_63 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_62); +lean_inc(x_63); +lean_inc(x_1); +x_64 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_64, 0, x_52); +lean_ctor_set(x_64, 1, x_53); +lean_ctor_set(x_64, 2, x_1); +lean_ctor_set(x_64, 3, x_54); +lean_ctor_set(x_64, 4, x_63); +lean_ctor_set(x_64, 5, x_55); +lean_ctor_set(x_64, 6, x_56); +lean_ctor_set(x_64, 7, x_57); +lean_ctor_set(x_64, 8, x_58); +lean_ctor_set(x_64, 9, x_59); +lean_ctor_set(x_64, 10, x_60); +lean_ctor_set_uint8(x_64, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_64, sizeof(void*)*11 + 1, x_2); +x_65 = lean_st_ref_get(x_9, x_10); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_ctor_get(x_66, 0); +lean_inc(x_68); +lean_dec(x_66); +x_69 = l_Lean_Kernel_isDiagnosticsEnabled(x_68); +lean_dec(x_68); +if (x_69 == 0) +{ +if (x_2 == 0) +{ +uint8_t x_90; +x_90 = 1; +x_70 = x_90; +goto block_89; +} +else +{ +uint8_t x_91; +x_91 = 0; +x_70 = x_91; +goto block_89; } } -LEAN_EXPORT lean_object* l_Lean_Meta_liftMetaM(lean_object* x_1, lean_object* x_2) { -_start: +else { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_liftMetaM___rarg), 2, 0); -return x_3; +if (x_2 == 0) +{ +uint8_t x_92; +x_92 = 0; +x_70 = x_92; +goto block_89; +} +else +{ +uint8_t x_93; +x_93 = 1; +x_70 = x_93; +goto block_89; } } -LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +block_89: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_apply_2(x_3, lean_box(0), x_1); -x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); -return x_10; +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_71 = lean_st_ref_take(x_9, x_67); +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, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_72, 1); +lean_inc(x_75); +x_76 = lean_ctor_get(x_72, 2); +lean_inc(x_76); +x_77 = lean_ctor_get(x_72, 3); +lean_inc(x_77); +x_78 = lean_ctor_get(x_72, 5); +lean_inc(x_78); +x_79 = lean_ctor_get(x_72, 6); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + lean_ctor_release(x_72, 2); + lean_ctor_release(x_72, 3); + lean_ctor_release(x_72, 4); + lean_ctor_release(x_72, 5); + lean_ctor_release(x_72, 6); + x_80 = x_72; +} else { + lean_dec_ref(x_72); + x_80 = lean_box(0); +} +x_81 = l_Lean_Kernel_enableDiag(x_74, x_2); +if (lean_is_scalar(x_80)) { + x_82 = lean_alloc_ctor(0, 7, 0); +} else { + x_82 = x_80; } +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_75); +lean_ctor_set(x_82, 2, x_76); +lean_ctor_set(x_82, 3, x_77); +lean_ctor_set(x_82, 4, x_6); +lean_ctor_set(x_82, 5, x_78); +lean_ctor_set(x_82, 6, x_79); +x_83 = lean_st_ref_set(x_9, x_82, x_73); +x_84 = lean_ctor_get(x_83, 1); +lean_inc(x_84); +lean_dec(x_83); +x_85 = lean_box(0); +x_86 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_63, x_2, x_3, x_4, x_5, x_85, x_64, x_9, x_84); +return x_86; } -LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg(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; -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_mapMetaM___rarg___lambda__1), 8, 2); -lean_closure_set(x_7, 0, x_5); -lean_closure_set(x_7, 1, x_3); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_apply_2(x_8, lean_box(0), x_7); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_apply_1(x_10, lean_box(0)); -x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); -return x_12; +lean_object* x_87; lean_object* x_88; +lean_dec(x_6); +x_87 = lean_box(0); +x_88 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_63, x_2, x_3, x_4, x_5, x_87, x_64, x_9, x_67); +return x_88; } } -LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM(lean_object* x_1) { +} +} +} +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__1() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_mapMetaM___rarg), 5, 0); -return x_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; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Context_config___default___closed__1; +x_3 = l_Lean_Meta_Context_lctx___default___closed__2; +x_4 = l_Lean_Meta_ParamInfo_backDeps___default___closed__1; +x_5 = lean_unsigned_to_nat(0u); +x_6 = 0; +x_7 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_7, 0, x_2); +lean_ctor_set(x_7, 1, x_3); +lean_ctor_set(x_7, 2, x_4); +lean_ctor_set(x_7, 3, x_1); +lean_ctor_set(x_7, 4, x_5); +lean_ctor_set(x_7, 5, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*6, x_6); +lean_ctor_set_uint8(x_7, sizeof(void*)*6 + 1, x_6); +return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___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) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2() { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_apply_1(x_1, x_3); -x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_State_mctx___default___closed__3; +x_3 = l_Lean_Meta_instInhabitedCache___closed__4; +x_4 = l_Lean_Meta_State_postponed___default___closed__3; +x_5 = l_Lean_Meta_instInhabitedDiagnostics___closed__1; +x_6 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_6, 0, x_2); +lean_ctor_set(x_6, 1, x_3); +lean_ctor_set(x_6, 2, x_1); +lean_ctor_set(x_6, 3, x_4); +lean_ctor_set(x_6, 4, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___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) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3() { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg___lambda__1), 8, 2); -lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_3); -x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg(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_Meta_instMetaEvalMetaM___rarg___closed__4() { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg___lambda__2), 8, 2); -lean_closure_set(x_7, 0, x_5); -lean_closure_set(x_7, 1, x_3); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_apply_2(x_8, lean_box(0), x_7); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_apply_1(x_10, lean_box(0)); -x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; +x_2 = l_Lean_Meta_ParamInfo_backDeps___default___closed__1; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM(lean_object* x_1, lean_object* x_2) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5() { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg), 5, 0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_firstFrontendMacroScope; +x_2 = lean_unsigned_to_nat(1u); +x_3 = lean_nat_add(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__6() { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_apply_2(x_1, x_3, x_4); -x_11 = lean_apply_7(x_2, lean_box(0), x_10, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_uniq", 5); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__7() { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg___lambda__1), 9, 2); -lean_closure_set(x_9, 0, x_1); -lean_closure_set(x_9, 1, x_3); -x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); -return x_10; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___rarg(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_Meta_instMetaEvalMetaM___rarg___closed__8() { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_6 = lean_ctor_get(x_2, 1); -lean_inc(x_6); -lean_dec(x_2); -x_7 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg___lambda__2), 8, 2); -lean_closure_set(x_7, 0, x_5); -lean_closure_set(x_7, 1, x_3); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = lean_apply_2(x_8, lean_box(0), x_7); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = lean_apply_1(x_10, lean_box(0)); -x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__7; +x_2 = lean_unsigned_to_nat(1u); +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_map2MetaM(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9() { _start: { -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg), 5, 0); +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; +x_1 = 1; +x_2 = l_Lean_Meta_State_mctx___default___closed__2; +x_3 = l_Lean_Meta_State_postponed___default___closed__3; +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Meta_instMetaEvalMetaM___rarg___closed__10() { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_7 = lean_st_ref_take(x_3, x_6); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 7); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___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) { +_start: +{ +lean_object* x_7; lean_object* x_21; lean_object* x_22; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; uint8_t x_73; +x_43 = lean_box(0); +x_44 = l_Lean_Core_getMaxHeartbeats(x_3); +x_45 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__5; +x_46 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__8; +x_47 = l_Lean_Meta_State_postponed___default___closed__3; +x_48 = l_Lean_Meta_instMonadEnvMetaM___lambda__2___closed__1; +x_49 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__9; +x_50 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_50, 0, x_2); +lean_ctor_set(x_50, 1, x_45); +lean_ctor_set(x_50, 2, x_46); +lean_ctor_set(x_50, 3, x_47); +lean_ctor_set(x_50, 4, x_48); +lean_ctor_set(x_50, 5, x_47); +lean_ctor_set(x_50, 6, x_49); +x_51 = lean_io_get_num_heartbeats(x_6); +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_instMetaEvalMetaM___rarg___closed__10; +x_55 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__4; +x_56 = lean_unsigned_to_nat(0u); +x_57 = lean_unsigned_to_nat(1000u); +x_58 = lean_box(0); +x_59 = lean_box(0); +x_60 = l_Lean_firstFrontendMacroScope; +x_61 = 0; +lean_inc(x_3); +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_54); +lean_ctor_set(x_62, 1, x_55); +lean_ctor_set(x_62, 2, x_3); +lean_ctor_set(x_62, 3, x_56); +lean_ctor_set(x_62, 4, x_57); +lean_ctor_set(x_62, 5, x_58); +lean_ctor_set(x_62, 6, x_59); +lean_ctor_set(x_62, 7, x_43); +lean_ctor_set(x_62, 8, x_52); +lean_ctor_set(x_62, 9, x_44); +lean_ctor_set(x_62, 10, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_61); +x_63 = lean_st_mk_ref(x_50, x_53); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = l_Lean_Meta_MetaM_toIO___rarg___closed__2; +x_67 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_3, x_66); +x_68 = lean_st_ref_get(x_64, x_65); +x_69 = lean_ctor_get(x_68, 0); +lean_inc(x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = lean_ctor_get(x_69, 0); +lean_inc(x_71); +lean_dec(x_69); +x_72 = l_Lean_Kernel_isDiagnosticsEnabled(x_71); +lean_dec(x_71); +if (x_72 == 0) +{ +if (x_67 == 0) +{ +uint8_t x_140; +x_140 = 1; +x_73 = x_140; +goto block_139; +} +else +{ +x_73 = x_61; +goto block_139; +} +} +else +{ +if (x_67 == 0) +{ +x_73 = x_61; +goto block_139; +} +else +{ +uint8_t x_141; +x_141 = 1; +x_73 = x_141; +goto block_139; +} +} +block_20: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; x_8 = lean_ctor_get(x_7, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_7, 1); lean_inc(x_9); lean_dec(x_7); -x_10 = !lean_is_exclusive(x_8); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); x_11 = lean_ctor_get(x_8, 1); -x_12 = lean_apply_1(x_1, x_11); -lean_ctor_set(x_8, 1, x_12); -x_13 = lean_st_ref_set(x_3, x_8, x_9); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = 1; +x_14 = lean_box(x_13); +x_15 = lean_apply_5(x_1, x_12, x_3, x_10, x_14, x_9); +return x_15; +} +else { -lean_object* x_15; lean_object* x_16; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -x_16 = lean_box(0); -lean_ctor_set(x_13, 0, x_16); -return x_13; +uint8_t x_16; +lean_dec(x_3); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_7); +if (x_16 == 0) +{ +return x_7; } else { lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_13, 1); +x_17 = lean_ctor_get(x_7, 0); +x_18 = lean_ctor_get(x_7, 1); +lean_inc(x_18); lean_inc(x_17); -lean_dec(x_13); -x_18 = lean_box(0); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); +lean_dec(x_7); +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; } } -else +} +block_42: { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_20 = lean_ctor_get(x_8, 0); -x_21 = lean_ctor_get(x_8, 1); -x_22 = lean_ctor_get(x_8, 2); -x_23 = lean_ctor_get(x_8, 3); -x_24 = lean_ctor_get(x_8, 4); -lean_inc(x_24); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_8); -x_25 = lean_apply_1(x_1, x_21); -x_26 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_26, 0, x_20); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_26, 2, x_22); -lean_ctor_set(x_26, 3, x_23); -lean_ctor_set(x_26, 4, x_24); -x_27 = lean_st_ref_set(x_3, x_26, x_9); -x_28 = lean_ctor_get(x_27, 1); +lean_dec(x_21); +x_24 = l_Lean_MessageData_toString(x_23, x_22); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set_tag(x_24, 1); +lean_ctor_set(x_24, 0, x_27); +x_7 = x_24; +goto block_20; +} +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); -if (lean_is_exclusive(x_27)) { - lean_ctor_release(x_27, 0); - lean_ctor_release(x_27, 1); - x_29 = x_27; -} else { - lean_dec_ref(x_27); - x_29 = lean_box(0); +lean_dec(x_24); +x_30 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_30, 0, x_28); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_7 = x_31; +goto block_20; } -x_30 = lean_box(0); -if (lean_is_scalar(x_29)) { - x_31 = lean_alloc_ctor(0, 2, 0); -} else { - x_31 = x_29; } -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_28); -return x_31; +else +{ +uint8_t x_32; +x_32 = !lean_is_exclusive(x_24); +if (x_32 == 0) +{ +x_7 = x_24; +goto block_20; } +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); +x_7 = x_35; +goto block_20; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +} +else { -lean_object* x_7; -x_7 = l_Lean_Meta_modifyCache(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_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_21, 0); +lean_inc(x_36); +lean_dec(x_21); +x_37 = l___private_Init_Data_Repr_0__Nat_reprFast(x_36); +x_38 = l_Lean_Meta_MetaM_toIO___rarg___closed__1; +x_39 = lean_string_append(x_38, x_37); +lean_dec(x_37); +x_40 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_40, 0, x_39); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_22); +x_7 = x_41; +goto block_20; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache(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: +block_139: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_st_ref_take(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) +if (x_73 == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_8, 1); -lean_dec(x_12); -x_13 = !lean_is_exclusive(x_9); -if (x_13 == 0) +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = lean_st_ref_take(x_64, x_70); +x_75 = lean_ctor_get(x_74, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = !lean_is_exclusive(x_75); +if (x_77 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_14 = lean_ctor_get(x_9, 0); -x_15 = lean_apply_1(x_1, x_14); -lean_ctor_set(x_9, 0, x_15); -x_16 = lean_st_ref_set(x_3, x_8, x_10); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_78 = lean_ctor_get(x_75, 0); +x_79 = lean_ctor_get(x_75, 4); +lean_dec(x_79); +x_80 = l_Lean_Kernel_enableDiag(x_78, x_67); +lean_ctor_set(x_75, 4, x_48); +lean_ctor_set(x_75, 0, x_80); +x_81 = lean_st_ref_set(x_64, x_75, x_76); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +x_83 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2; +x_84 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__1; +x_85 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_86 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(x_3, x_67, x_83, x_4, x_84, x_48, x_85, x_62, x_64, x_82); +if (lean_obj_tag(x_86) == 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; +lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_87 = lean_ctor_get(x_86, 0); +lean_inc(x_87); +x_88 = lean_ctor_get(x_86, 1); +lean_inc(x_88); +lean_dec(x_86); +x_89 = lean_st_ref_get(x_64, x_88); +lean_dec(x_64); +x_90 = !lean_is_exclusive(x_89); +if (x_90 == 0) +{ +lean_object* x_91; lean_object* x_92; +x_91 = lean_ctor_get(x_89, 0); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_87); +lean_ctor_set(x_92, 1, x_91); +lean_ctor_set(x_89, 0, x_92); +x_7 = x_89; +goto block_20; } else { -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; +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_93 = lean_ctor_get(x_89, 0); +x_94 = lean_ctor_get(x_89, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_89); +x_95 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_95, 0, x_87); +lean_ctor_set(x_95, 1, 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_94); +x_7 = x_96; +goto block_20; } } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_23 = lean_ctor_get(x_9, 0); -x_24 = lean_ctor_get(x_9, 1); -x_25 = lean_ctor_get(x_9, 2); -x_26 = lean_ctor_get(x_9, 3); -x_27 = lean_ctor_get(x_9, 4); -x_28 = lean_ctor_get(x_9, 5); -x_29 = lean_ctor_get(x_9, 6); -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_dec(x_9); -x_30 = lean_apply_1(x_1, x_23); -x_31 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_24); -lean_ctor_set(x_31, 2, x_25); -lean_ctor_set(x_31, 3, x_26); -lean_ctor_set(x_31, 4, x_27); -lean_ctor_set(x_31, 5, x_28); -lean_ctor_set(x_31, 6, x_29); -lean_ctor_set(x_8, 1, x_31); -x_32 = lean_st_ref_set(x_3, x_8, x_10); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_34 = x_32; +lean_object* x_97; lean_object* x_98; +lean_dec(x_64); +x_97 = lean_ctor_get(x_86, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_86, 1); +lean_inc(x_98); +lean_dec(x_86); +x_21 = x_97; +x_22 = x_98; +goto block_42; +} +} +else +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_99 = lean_ctor_get(x_75, 0); +x_100 = lean_ctor_get(x_75, 1); +x_101 = lean_ctor_get(x_75, 2); +x_102 = lean_ctor_get(x_75, 3); +x_103 = lean_ctor_get(x_75, 5); +x_104 = lean_ctor_get(x_75, 6); +lean_inc(x_104); +lean_inc(x_103); +lean_inc(x_102); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_75); +x_105 = l_Lean_Kernel_enableDiag(x_99, x_67); +x_106 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_100); +lean_ctor_set(x_106, 2, x_101); +lean_ctor_set(x_106, 3, x_102); +lean_ctor_set(x_106, 4, x_48); +lean_ctor_set(x_106, 5, x_103); +lean_ctor_set(x_106, 6, x_104); +x_107 = lean_st_ref_set(x_64, x_106, x_76); +x_108 = lean_ctor_get(x_107, 1); +lean_inc(x_108); +lean_dec(x_107); +x_109 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2; +x_110 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__1; +x_111 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_112 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(x_3, x_67, x_109, x_4, x_110, x_48, x_111, x_62, x_64, x_108); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +x_115 = lean_st_ref_get(x_64, x_114); +lean_dec(x_64); +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_32); - x_34 = lean_box(0); + lean_dec_ref(x_115); + x_118 = lean_box(0); } -x_35 = lean_box(0); -if (lean_is_scalar(x_34)) { - x_36 = lean_alloc_ctor(0, 2, 0); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_113); +lean_ctor_set(x_119, 1, x_116); +if (lean_is_scalar(x_118)) { + x_120 = lean_alloc_ctor(0, 2, 0); } else { - x_36 = x_34; + x_120 = x_118; +} +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +x_7 = x_120; +goto block_20; +} +else +{ +lean_object* x_121; lean_object* x_122; +lean_dec(x_64); +x_121 = lean_ctor_get(x_112, 0); +lean_inc(x_121); +x_122 = lean_ctor_get(x_112, 1); +lean_inc(x_122); +lean_dec(x_112); +x_21 = x_121; +x_22 = x_122; +goto block_42; } -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_33); -return x_36; } } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); -x_38 = lean_ctor_get(x_8, 2); -x_39 = lean_ctor_get(x_8, 3); -x_40 = lean_ctor_get(x_8, 4); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_8); -x_41 = lean_ctor_get(x_9, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_9, 1); -lean_inc(x_42); -x_43 = lean_ctor_get(x_9, 2); -lean_inc(x_43); -x_44 = lean_ctor_get(x_9, 3); -lean_inc(x_44); -x_45 = lean_ctor_get(x_9, 4); -lean_inc(x_45); -x_46 = lean_ctor_get(x_9, 5); -lean_inc(x_46); -x_47 = lean_ctor_get(x_9, 6); -lean_inc(x_47); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - lean_ctor_release(x_9, 2); - lean_ctor_release(x_9, 3); - lean_ctor_release(x_9, 4); - lean_ctor_release(x_9, 5); - lean_ctor_release(x_9, 6); - x_48 = x_9; -} else { - lean_dec_ref(x_9); - x_48 = lean_box(0); +lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_123 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__2; +x_124 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__1; +x_125 = lean_box(0); +lean_inc(x_64); +lean_inc(x_3); +x_126 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(x_3, x_67, x_123, x_4, x_124, x_48, x_125, x_62, x_64, x_70); +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; uint8_t x_130; +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_st_ref_get(x_64, x_128); +lean_dec(x_64); +x_130 = !lean_is_exclusive(x_129); +if (x_130 == 0) +{ +lean_object* x_131; lean_object* x_132; +x_131 = lean_ctor_get(x_129, 0); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_127); +lean_ctor_set(x_132, 1, x_131); +lean_ctor_set(x_129, 0, x_132); +x_7 = x_129; +goto block_20; } -x_49 = lean_apply_1(x_1, x_41); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 7, 0); -} else { - x_50 = x_48; +else +{ +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_133 = lean_ctor_get(x_129, 0); +x_134 = lean_ctor_get(x_129, 1); +lean_inc(x_134); +lean_inc(x_133); +lean_dec(x_129); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_127); +lean_ctor_set(x_135, 1, x_133); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_135); +lean_ctor_set(x_136, 1, x_134); +x_7 = x_136; +goto block_20; } -lean_ctor_set(x_50, 0, x_49); -lean_ctor_set(x_50, 1, x_42); -lean_ctor_set(x_50, 2, x_43); -lean_ctor_set(x_50, 3, x_44); -lean_ctor_set(x_50, 4, x_45); -lean_ctor_set(x_50, 5, x_46); -lean_ctor_set(x_50, 6, x_47); -x_51 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_51, 0, x_37); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_38); -lean_ctor_set(x_51, 3, x_39); -lean_ctor_set(x_51, 4, x_40); -x_52 = lean_st_ref_set(x_3, x_51, x_10); -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; +else +{ +lean_object* x_137; lean_object* x_138; +lean_dec(x_64); +x_137 = lean_ctor_get(x_126, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_126, 1); +lean_inc(x_138); +lean_dec(x_126); +x_21 = x_137; +x_22 = x_138; +goto block_42; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache___boxed(lean_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_instMetaEvalMetaM(lean_object* x_1) { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_modifyInferTypeCache(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_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_instMetaEvalMetaM___rarg___boxed), 6, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache(lean_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_instMetaEvalMetaM___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) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_st_ref_take(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_3); +lean_dec(x_3); +x_12 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__1(x_1, x_2, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_7); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_8, 1); -lean_dec(x_12); -x_13 = !lean_is_exclusive(x_9); -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_9, 5); -x_15 = lean_apply_1(x_1, x_14); -lean_ctor_set(x_9, 5, x_15); -x_16 = lean_st_ref_set(x_3, x_8, x_10); -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; +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___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) { +_start: { -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; +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_Meta_instMetaEvalMetaM___rarg___lambda__2(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_instMetaEvalMetaM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_23 = lean_ctor_get(x_9, 0); -x_24 = lean_ctor_get(x_9, 1); -x_25 = lean_ctor_get(x_9, 2); -x_26 = lean_ctor_get(x_9, 3); -x_27 = lean_ctor_get(x_9, 4); -x_28 = lean_ctor_get(x_9, 5); -x_29 = lean_ctor_get(x_9, 6); -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_dec(x_9); -x_30 = lean_apply_1(x_1, x_28); -x_31 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_31, 0, x_23); -lean_ctor_set(x_31, 1, x_24); -lean_ctor_set(x_31, 2, x_25); -lean_ctor_set(x_31, 3, x_26); -lean_ctor_set(x_31, 4, x_27); -lean_ctor_set(x_31, 5, x_30); -lean_ctor_set(x_31, 6, x_29); -lean_ctor_set(x_8, 1, x_31); -x_32 = lean_st_ref_set(x_3, x_8, x_10); -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); +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_5); +lean_dec(x_5); +x_8 = l_Lean_Meta_instMetaEvalMetaM___rarg(x_1, x_2, x_3, x_4, x_7, x_6); +return x_8; } -x_35 = lean_box(0); -if (lean_is_scalar(x_34)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_34; } -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_33); -return x_36; +static lean_object* _init_l_Lean_Meta_throwIsDefEqStuck___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_isDefEqStuckExceptionId; +return x_1; } } -else +static lean_object* _init_l_Lean_Meta_throwIsDefEqStuck___rarg___closed__2() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); -x_38 = lean_ctor_get(x_8, 2); -x_39 = lean_ctor_get(x_8, 3); -x_40 = lean_ctor_get(x_8, 4); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_8); -x_41 = lean_ctor_get(x_9, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_9, 1); -lean_inc(x_42); -x_43 = lean_ctor_get(x_9, 2); -lean_inc(x_43); -x_44 = lean_ctor_get(x_9, 3); -lean_inc(x_44); -x_45 = lean_ctor_get(x_9, 4); -lean_inc(x_45); -x_46 = lean_ctor_get(x_9, 5); -lean_inc(x_46); -x_47 = lean_ctor_get(x_9, 6); -lean_inc(x_47); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - lean_ctor_release(x_9, 2); - lean_ctor_release(x_9, 3); - lean_ctor_release(x_9, 4); - lean_ctor_release(x_9, 5); - lean_ctor_release(x_9, 6); - x_48 = x_9; -} else { - lean_dec_ref(x_9); - x_48 = lean_box(0); -} -x_49 = lean_apply_1(x_1, x_46); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 7, 0); -} else { - x_50 = x_48; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_throwIsDefEqStuck___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_ctor_set(x_50, 0, x_41); -lean_ctor_set(x_50, 1, x_42); -lean_ctor_set(x_50, 2, x_43); -lean_ctor_set(x_50, 3, x_44); -lean_ctor_set(x_50, 4, x_45); -lean_ctor_set(x_50, 5, x_49); -lean_ctor_set(x_50, 6, x_47); -x_51 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_51, 0, x_37); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_38); -lean_ctor_set(x_51, 3, x_39); -lean_ctor_set(x_51, 4, x_40); -x_52 = lean_st_ref_set(x_3, x_51, x_10); -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_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck___rarg(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Meta_throwIsDefEqStuck___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_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; } +LEAN_EXPORT lean_object* l_Lean_Meta_throwIsDefEqStuck(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; +x_6 = lean_alloc_closure((void*)(l_Lean_Meta_throwIsDefEqStuck___rarg), 1, 0); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache___boxed(lean_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_throwIsDefEqStuck___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_7; -x_7 = l_Lean_Meta_modifyDefEqTransientCache(x_1, x_2, x_3, x_4, x_5, x_6); +lean_object* x_6; +x_6 = l_Lean_Meta_throwIsDefEqStuck(x_1, x_2, x_3, x_4, x_5); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_7; +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1() { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_7 = lean_st_ref_take(x_3, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_8, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_7, 1); -lean_inc(x_10); -lean_dec(x_7); -x_11 = !lean_is_exclusive(x_8); -if (x_11 == 0) -{ -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_8, 1); -lean_dec(x_12); -x_13 = !lean_is_exclusive(x_9); -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_9, 6); -x_15 = lean_apply_1(x_1, x_14); -lean_ctor_set(x_9, 6, x_15); -x_16 = lean_st_ref_set(x_3, x_8, x_10); -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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Meta", 4); +return x_1; } -else +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__2() { +_start: { -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; +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_Basic___hyg_1758____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__3() { +_start: { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_23 = lean_ctor_get(x_9, 0); -x_24 = lean_ctor_get(x_9, 1); -x_25 = lean_ctor_get(x_9, 2); -x_26 = lean_ctor_get(x_9, 3); -x_27 = lean_ctor_get(x_9, 4); -x_28 = lean_ctor_get(x_9, 5); -x_29 = lean_ctor_get(x_9, 6); -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_dec(x_9); -x_30 = lean_apply_1(x_1, x_29); -x_31 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_31, 0, x_23); -lean_ctor_set(x_31, 1, x_24); -lean_ctor_set(x_31, 2, x_25); -lean_ctor_set(x_31, 3, x_26); -lean_ctor_set(x_31, 4, x_27); -lean_ctor_set(x_31, 5, x_28); -lean_ctor_set(x_31, 6, x_30); -lean_ctor_set(x_8, 1, x_31); -x_32 = lean_st_ref_set(x_3, x_8, x_10); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_34 = x_32; -} else { - lean_dec_ref(x_32); - x_34 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean", 4); +return x_1; } -x_35 = lean_box(0); -if (lean_is_scalar(x_34)) { - x_36 = lean_alloc_ctor(0, 2, 0); -} else { - x_36 = x_34; } -lean_ctor_set(x_36, 0, x_35); -lean_ctor_set(x_36, 1, x_33); -return x_36; +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__4() { +_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_Basic___hyg_1758____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__5() { +_start: { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); -x_38 = lean_ctor_get(x_8, 2); -x_39 = lean_ctor_get(x_8, 3); -x_40 = lean_ctor_get(x_8, 4); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_8); -x_41 = lean_ctor_get(x_9, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_9, 1); -lean_inc(x_42); -x_43 = lean_ctor_get(x_9, 2); -lean_inc(x_43); -x_44 = lean_ctor_get(x_9, 3); -lean_inc(x_44); -x_45 = lean_ctor_get(x_9, 4); -lean_inc(x_45); -x_46 = lean_ctor_get(x_9, 5); -lean_inc(x_46); -x_47 = lean_ctor_get(x_9, 6); -lean_inc(x_47); -if (lean_is_exclusive(x_9)) { - lean_ctor_release(x_9, 0); - lean_ctor_release(x_9, 1); - lean_ctor_release(x_9, 2); - lean_ctor_release(x_9, 3); - lean_ctor_release(x_9, 4); - lean_ctor_release(x_9, 5); - lean_ctor_release(x_9, 6); - x_48 = x_9; -} else { - lean_dec_ref(x_9); - x_48 = lean_box(0); -} -x_49 = lean_apply_1(x_1, x_47); -if (lean_is_scalar(x_48)) { - x_50 = lean_alloc_ctor(0, 7, 0); -} else { - x_50 = x_48; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_50, 0, x_41); -lean_ctor_set(x_50, 1, x_42); -lean_ctor_set(x_50, 2, x_43); -lean_ctor_set(x_50, 3, x_44); -lean_ctor_set(x_50, 4, x_45); -lean_ctor_set(x_50, 5, x_46); -lean_ctor_set(x_50, 6, x_49); -x_51 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_51, 0, x_37); -lean_ctor_set(x_51, 1, x_50); -lean_ctor_set(x_51, 2, x_38); -lean_ctor_set(x_51, 3, x_39); -lean_ctor_set(x_51, 4, x_40); -x_52 = lean_st_ref_set(x_3, x_51, x_10); -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; +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("initFn", 6); +return x_1; } -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_53); -return x_56; } +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__6; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__8() { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_modifyDefEqPermCache(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_@", 2); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__9() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = lean_st_ref_take(x_1, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -x_8 = lean_ctor_get(x_5, 1); -lean_inc(x_8); -lean_dec(x_5); -x_9 = !lean_is_exclusive(x_6); -if (x_9 == 0) -{ -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_6, 1); -lean_dec(x_10); -x_11 = !lean_is_exclusive(x_7); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_12 = lean_ctor_get(x_7, 6); -lean_dec(x_12); -x_13 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; -lean_ctor_set(x_7, 6, x_13); -x_14 = lean_st_ref_set(x_1, x_6, x_8); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__7; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__8; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__10() { +_start: { -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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__9; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__11() { +_start: { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_21 = lean_ctor_get(x_7, 0); -x_22 = lean_ctor_get(x_7, 1); -x_23 = lean_ctor_get(x_7, 2); -x_24 = lean_ctor_get(x_7, 3); -x_25 = lean_ctor_get(x_7, 4); -x_26 = lean_ctor_get(x_7, 5); -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_dec(x_7); -x_27 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; -x_28 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_28, 0, x_21); -lean_ctor_set(x_28, 1, x_22); -lean_ctor_set(x_28, 2, x_23); -lean_ctor_set(x_28, 3, x_24); -lean_ctor_set(x_28, 4, x_25); -lean_ctor_set(x_28, 5, x_26); -lean_ctor_set(x_28, 6, x_27); -lean_ctor_set(x_6, 1, x_28); -x_29 = lean_st_ref_set(x_1, x_6, x_8); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); -if (lean_is_exclusive(x_29)) { - lean_ctor_release(x_29, 0); - lean_ctor_release(x_29, 1); - x_31 = x_29; -} else { - lean_dec_ref(x_29); - x_31 = lean_box(0); -} -x_32 = lean_box(0); -if (lean_is_scalar(x_31)) { - x_33 = lean_alloc_ctor(0, 2, 0); -} else { - x_33 = x_31; -} -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_30); -return x_33; -} -} -else -{ -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_34 = lean_ctor_get(x_6, 0); -x_35 = lean_ctor_get(x_6, 2); -x_36 = lean_ctor_get(x_6, 3); -x_37 = lean_ctor_get(x_6, 4); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_6); -x_38 = lean_ctor_get(x_7, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_7, 1); -lean_inc(x_39); -x_40 = lean_ctor_get(x_7, 2); -lean_inc(x_40); -x_41 = lean_ctor_get(x_7, 3); -lean_inc(x_41); -x_42 = lean_ctor_get(x_7, 4); -lean_inc(x_42); -x_43 = lean_ctor_get(x_7, 5); -lean_inc(x_43); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - x_44 = x_7; -} else { - lean_dec_ref(x_7); - x_44 = lean_box(0); -} -x_45 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; -if (lean_is_scalar(x_44)) { - x_46 = lean_alloc_ctor(0, 7, 0); -} else { - x_46 = x_44; -} -lean_ctor_set(x_46, 0, x_38); -lean_ctor_set(x_46, 1, x_39); -lean_ctor_set(x_46, 2, x_40); -lean_ctor_set(x_46, 3, x_41); -lean_ctor_set(x_46, 4, x_42); -lean_ctor_set(x_46, 5, x_43); -lean_ctor_set(x_46, 6, x_45); -x_47 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_47, 0, x_34); -lean_ctor_set(x_47, 1, x_46); -lean_ctor_set(x_47, 2, x_35); -lean_ctor_set(x_47, 3, x_36); -lean_ctor_set(x_47, 4, x_37); -x_48 = lean_st_ref_set(x_1, x_47, x_8); -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; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_resetDefEqPermCaches___rarg___boxed), 4, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_Meta_resetDefEqPermCaches___rarg(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l_Lean_Meta_resetDefEqPermCaches(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDiag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_1); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_7, 0); -lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_16 = lean_ctor_get(x_7, 1); -lean_inc(x_16); -lean_dec(x_7); -x_17 = lean_st_ref_take(x_3, 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_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_21 = lean_ctor_get(x_18, 4); -x_22 = lean_apply_1(x_1, x_21); -lean_ctor_set(x_18, 4, x_22); -x_23 = lean_st_ref_set(x_3, x_18, x_19); -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_box(0); -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_box(0); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -else -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_30 = lean_ctor_get(x_18, 0); -x_31 = lean_ctor_get(x_18, 1); -x_32 = lean_ctor_get(x_18, 2); -x_33 = lean_ctor_get(x_18, 3); -x_34 = lean_ctor_get(x_18, 4); -lean_inc(x_34); -lean_inc(x_33); -lean_inc(x_32); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_18); -x_35 = lean_apply_1(x_1, x_34); -x_36 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_36, 0, x_30); -lean_ctor_set(x_36, 1, x_31); -lean_ctor_set(x_36, 2, x_32); -lean_ctor_set(x_36, 3, x_33); -lean_ctor_set(x_36, 4, x_35); -x_37 = lean_st_ref_set(x_3, x_36, x_19); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -if (lean_is_exclusive(x_37)) { - lean_ctor_release(x_37, 0); - lean_ctor_release(x_37, 1); - x_39 = x_37; -} else { - lean_dec_ref(x_37); - x_39 = lean_box(0); -} -x_40 = lean_box(0); -if (lean_is_scalar(x_39)) { - x_41 = lean_alloc_ctor(0, 2, 0); -} else { - x_41 = x_39; -} -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_38); -return x_41; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_modifyDiag___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_modifyDiag(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_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; uint8_t x_7; -x_6 = lean_array_get_size(x_1); -x_7 = lean_nat_dec_lt(x_4, x_6); -lean_dec(x_6); -if (x_7 == 0) -{ -lean_object* x_8; -lean_dec(x_4); -x_8 = lean_box(0); -return x_8; -} -else -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_array_fget(x_1, x_4); -x_10 = lean_name_eq(x_5, x_9); -lean_dec(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_4, x_11); -lean_dec(x_4); -x_3 = lean_box(0); -x_4 = x_12; -goto _start; -} -else -{ -lean_object* x_14; lean_object* x_15; -x_14 = lean_array_fget(x_2, x_4); -lean_dec(x_4); -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; -} -} -} -} -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__1() { -_start: -{ -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = 5; -x_3 = lean_usize_shift_left(x_1, x_2); -return x_3; -} -} -static size_t _init_l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__2() { -_start: -{ -size_t x_1; size_t x_2; size_t x_3; -x_1 = 1; -x_2 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__1; -x_3 = lean_usize_sub(x_2, x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2(lean_object* x_1, size_t x_2, lean_object* x_3) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_4; size_t x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = 5; -x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__2; -x_7 = lean_usize_land(x_2, x_6); -x_8 = lean_usize_to_nat(x_7); -x_9 = lean_box(2); -x_10 = lean_array_get(x_9, x_4, x_8); -lean_dec(x_8); -lean_dec(x_4); -switch (lean_obj_tag(x_10)) { -case 0: -{ -lean_object* x_11; lean_object* x_12; uint8_t 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_name_eq(x_3, x_11); -lean_dec(x_11); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_12); -x_14 = lean_box(0); -return x_14; -} -else -{ -lean_object* x_15; -x_15 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_15, 0, x_12); -return x_15; -} -} -case 1: -{ -lean_object* x_16; size_t x_17; -x_16 = lean_ctor_get(x_10, 0); -lean_inc(x_16); -lean_dec(x_10); -x_17 = lean_usize_shift_right(x_2, x_5); -x_1 = x_16; -x_2 = x_17; -goto _start; -} -default: -{ -lean_object* x_19; -x_19 = lean_box(0); -return x_19; -} -} -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_1, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_1, 1); -lean_inc(x_21); -lean_dec(x_1); -x_22 = lean_unsigned_to_nat(0u); -x_23 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3(x_20, x_21, lean_box(0), x_22, x_3); -lean_dec(x_21); -lean_dec(x_20); -return x_23; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; uint64_t x_4; size_t x_5; lean_object* x_6; -x_3 = lean_ctor_get(x_1, 0); -lean_inc(x_3); -lean_dec(x_1); -x_4 = l_Lean_Name_hash___override(x_2); -x_5 = lean_uint64_to_usize(x_4); -x_6 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2(x_3, x_5, x_2); -lean_dec(x_2); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6(size_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -lean_dec(x_5); -return x_6; -} -else -{ -lean_object* x_9; lean_object* x_10; uint64_t x_11; size_t x_12; size_t x_13; size_t x_14; size_t x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_9 = lean_array_fget(x_2, x_5); -x_10 = lean_array_fget(x_3, x_5); -x_11 = l_Lean_Name_hash___override(x_9); -x_12 = lean_uint64_to_usize(x_11); -x_13 = 1; -x_14 = lean_usize_sub(x_1, x_13); -x_15 = 5; -x_16 = lean_usize_mul(x_15, x_14); -x_17 = lean_usize_shift_right(x_12, x_16); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_add(x_5, x_18); -lean_dec(x_5); -x_20 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_6, x_17, x_1, x_9, x_10); -x_4 = lean_box(0); -x_5 = x_19; -x_6 = x_20; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_recordUnfold___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_array_get_size(x_5); -x_8 = lean_nat_dec_lt(x_2, x_7); -lean_dec(x_7); -if (x_8 == 0) -{ -uint8_t x_9; -lean_dec(x_2); -x_9 = !lean_is_exclusive(x_1); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_1, 1); -lean_dec(x_10); -x_11 = lean_ctor_get(x_1, 0); -lean_dec(x_11); -x_12 = lean_array_push(x_5, x_3); -x_13 = lean_array_push(x_6, x_4); -lean_ctor_set(x_1, 1, x_13); -lean_ctor_set(x_1, 0, x_12); -return x_1; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_1); -x_14 = lean_array_push(x_5, x_3); -x_15 = lean_array_push(x_6, x_4); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; -} -} -else -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_array_fget(x_5, x_2); -x_18 = lean_name_eq(x_3, x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_6); -lean_dec(x_5); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_2, x_19); -lean_dec(x_2); -x_2 = x_20; -goto _start; -} -else -{ -uint8_t x_22; -x_22 = !lean_is_exclusive(x_1); -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_1, 1); -lean_dec(x_23); -x_24 = lean_ctor_get(x_1, 0); -lean_dec(x_24); -x_25 = lean_array_fset(x_5, x_2, x_3); -x_26 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -lean_ctor_set(x_1, 1, x_26); -lean_ctor_set(x_1, 0, x_25); -return x_1; -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_1); -x_27 = lean_array_fset(x_5, x_2, x_3); -x_28 = lean_array_fset(x_6, x_2, x_4); -lean_dec(x_2); -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; -} -} -} -} -} -static lean_object* _init_l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_mkEmptyEntries(lean_box(0), lean_box(0)); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -uint8_t x_6; -x_6 = !lean_is_exclusive(x_1); -if (x_6 == 0) -{ -lean_object* x_7; size_t x_8; size_t x_9; size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = lean_ctor_get(x_1, 0); -x_8 = 1; -x_9 = 5; -x_10 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__2; -x_11 = lean_usize_land(x_2, x_10); -x_12 = lean_usize_to_nat(x_11); -x_13 = lean_array_get_size(x_7); -x_14 = lean_nat_dec_lt(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -lean_dec(x_12); -lean_dec(x_5); -lean_dec(x_4); -return x_1; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_array_fget(x_7, x_12); -x_16 = lean_box(0); -x_17 = lean_array_fset(x_7, x_12, x_16); -switch (lean_obj_tag(x_15)) { -case 0: -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -x_21 = lean_name_eq(x_4, x_19); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -lean_free_object(x_15); -x_22 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_19, x_20, x_4, x_5); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_array_fset(x_17, x_12, x_23); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_24); -return x_1; -} -else -{ -lean_object* x_25; -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_25 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_25); -return x_1; -} -} -else -{ -lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_26 = lean_ctor_get(x_15, 0); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -lean_inc(x_26); -lean_dec(x_15); -x_28 = lean_name_eq(x_4, x_26); -if (x_28 == 0) -{ -lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_29 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_26, x_27, 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; -} -else -{ -lean_object* x_32; lean_object* x_33; -lean_dec(x_27); -lean_dec(x_26); -x_32 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_32, 0, x_4); -lean_ctor_set(x_32, 1, x_5); -x_33 = lean_array_fset(x_17, x_12, x_32); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_33); -return x_1; -} -} -} -case 1: -{ -uint8_t x_34; -x_34 = !lean_is_exclusive(x_15); -if (x_34 == 0) -{ -lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; -x_35 = lean_ctor_get(x_15, 0); -x_36 = lean_usize_shift_right(x_2, x_9); -x_37 = lean_usize_add(x_3, x_8); -x_38 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_35, x_36, x_37, x_4, x_5); -lean_ctor_set(x_15, 0, x_38); -x_39 = lean_array_fset(x_17, x_12, x_15); -lean_dec(x_12); -lean_ctor_set(x_1, 0, x_39); -return x_1; -} -else -{ -lean_object* x_40; size_t x_41; size_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_40 = lean_ctor_get(x_15, 0); -lean_inc(x_40); -lean_dec(x_15); -x_41 = lean_usize_shift_right(x_2, x_9); -x_42 = lean_usize_add(x_3, x_8); -x_43 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_40, x_41, x_42, 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_dec(x_12); -lean_ctor_set(x_1, 0, x_45); -return x_1; -} -} -default: -{ -lean_object* x_46; lean_object* x_47; -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_dec(x_12); -lean_ctor_set(x_1, 0, x_47); -return x_1; -} -} -} -} -else -{ -lean_object* x_48; size_t x_49; size_t x_50; size_t x_51; size_t x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_48 = lean_ctor_get(x_1, 0); -lean_inc(x_48); -lean_dec(x_1); -x_49 = 1; -x_50 = 5; -x_51 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___closed__2; -x_52 = lean_usize_land(x_2, x_51); -x_53 = lean_usize_to_nat(x_52); -x_54 = lean_array_get_size(x_48); -x_55 = lean_nat_dec_lt(x_53, x_54); -lean_dec(x_54); -if (x_55 == 0) -{ -lean_object* x_56; -lean_dec(x_53); -lean_dec(x_5); -lean_dec(x_4); -x_56 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_56, 0, x_48); -return x_56; -} -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_array_fget(x_48, x_53); -x_58 = lean_box(0); -x_59 = lean_array_fset(x_48, x_53, x_58); -switch (lean_obj_tag(x_57)) { -case 0: -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t x_63; -x_60 = lean_ctor_get(x_57, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_57, 1); -lean_inc(x_61); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_62 = x_57; -} else { - lean_dec_ref(x_57); - x_62 = lean_box(0); -} -x_63 = lean_name_eq(x_4, x_60); -if (x_63 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_62); -x_64 = l_Lean_PersistentHashMap_mkCollisionNode___rarg(x_60, x_61, x_4, x_5); -x_65 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_65, 0, x_64); -x_66 = lean_array_fset(x_59, x_53, x_65); -lean_dec(x_53); -x_67 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_67, 0, x_66); -return x_67; -} -else -{ -lean_object* x_68; lean_object* x_69; lean_object* x_70; -lean_dec(x_61); -lean_dec(x_60); -if (lean_is_scalar(x_62)) { - x_68 = lean_alloc_ctor(0, 2, 0); -} else { - x_68 = x_62; -} -lean_ctor_set(x_68, 0, x_4); -lean_ctor_set(x_68, 1, x_5); -x_69 = lean_array_fset(x_59, x_53, x_68); -lean_dec(x_53); -x_70 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_70, 0, x_69); -return x_70; -} -} -case 1: -{ -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; -x_71 = lean_ctor_get(x_57, 0); -lean_inc(x_71); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - x_72 = x_57; -} else { - lean_dec_ref(x_57); - x_72 = lean_box(0); -} -x_73 = lean_usize_shift_right(x_2, x_50); -x_74 = lean_usize_add(x_3, x_49); -x_75 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_71, x_73, x_74, x_4, x_5); -if (lean_is_scalar(x_72)) { - x_76 = lean_alloc_ctor(1, 1, 0); -} else { - x_76 = x_72; -} -lean_ctor_set(x_76, 0, x_75); -x_77 = lean_array_fset(x_59, x_53, x_76); -lean_dec(x_53); -x_78 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_78, 0, x_77); -return x_78; -} -default: -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_4); -lean_ctor_set(x_79, 1, x_5); -x_80 = lean_array_fset(x_59, x_53, x_79); -lean_dec(x_53); -x_81 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_81, 0, x_80); -return x_81; -} -} -} -} -} -else -{ -uint8_t x_82; -x_82 = !lean_is_exclusive(x_1); -if (x_82 == 0) -{ -lean_object* x_83; lean_object* x_84; size_t x_85; uint8_t x_86; -x_83 = lean_unsigned_to_nat(0u); -x_84 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_recordUnfold___spec__7(x_1, x_83, x_4, x_5); -x_85 = 7; -x_86 = lean_usize_dec_le(x_85, x_3); -if (x_86 == 0) -{ -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_84); -x_88 = lean_unsigned_to_nat(4u); -x_89 = lean_nat_dec_lt(x_87, x_88); -lean_dec(x_87); -if (x_89 == 0) -{ -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_90 = lean_ctor_get(x_84, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_84, 1); -lean_inc(x_91); -lean_dec(x_84); -x_92 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___closed__1; -x_93 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6(x_3, x_90, x_91, lean_box(0), x_83, x_92); -lean_dec(x_91); -lean_dec(x_90); -return x_93; -} -else -{ -return x_84; -} -} -else -{ -return x_84; -} -} -else -{ -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; size_t x_99; uint8_t x_100; -x_94 = lean_ctor_get(x_1, 0); -x_95 = lean_ctor_get(x_1, 1); -lean_inc(x_95); -lean_inc(x_94); -lean_dec(x_1); -x_96 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_96, 0, x_94); -lean_ctor_set(x_96, 1, x_95); -x_97 = lean_unsigned_to_nat(0u); -x_98 = l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_recordUnfold___spec__7(x_96, x_97, x_4, x_5); -x_99 = 7; -x_100 = lean_usize_dec_le(x_99, x_3); -if (x_100 == 0) -{ -lean_object* x_101; lean_object* x_102; uint8_t x_103; -x_101 = l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(x_98); -x_102 = lean_unsigned_to_nat(4u); -x_103 = lean_nat_dec_lt(x_101, x_102); -lean_dec(x_101); -if (x_103 == 0) -{ -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_ctor_get(x_98, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_98, 1); -lean_inc(x_105); -lean_dec(x_98); -x_106 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___closed__1; -x_107 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6(x_3, x_104, x_105, lean_box(0), x_97, x_106); -lean_dec(x_105); -lean_dec(x_104); -return x_107; -} -else -{ -return x_98; -} -} -else -{ -return x_98; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; uint64_t x_7; size_t x_8; size_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_Name_hash___override(x_2); -x_8 = lean_uint64_to_usize(x_7); -x_9 = 1; -x_10 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_5, x_8, x_9, x_2, x_3); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_6, x_11); -lean_dec(x_6); -lean_ctor_set(x_1, 1, x_12); -lean_ctor_set(x_1, 0, x_10); -return x_1; -} -else -{ -lean_object* x_13; lean_object* x_14; uint64_t 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; -x_13 = lean_ctor_get(x_1, 0); -x_14 = lean_ctor_get(x_1, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_1); -x_15 = l_Lean_Name_hash___override(x_2); -x_16 = lean_uint64_to_usize(x_15); -x_17 = 1; -x_18 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_13, x_16, x_17, x_2, x_3); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_add(x_14, x_19); -lean_dec(x_14); -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; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_1); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_7, 0); -lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_16 = lean_ctor_get(x_7, 1); -lean_inc(x_16); -lean_dec(x_7); -x_17 = lean_st_ref_take(x_3, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 4); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_18, 4); -lean_dec(x_22); -x_23 = !lean_is_exclusive(x_19); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_19, 0); -lean_inc(x_1); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_24, x_1); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_26); -lean_ctor_set(x_19, 0, x_27); -x_28 = lean_st_ref_set(x_3, x_18, x_20); -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; uint8_t x_40; -x_35 = lean_ctor_get(x_25, 0); -lean_inc(x_35); -lean_dec(x_25); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_35, x_36); -lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_37); -lean_ctor_set(x_19, 0, x_38); -x_39 = lean_st_ref_set(x_3, x_18, x_20); -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; lean_object* x_49; -x_46 = lean_ctor_get(x_19, 0); -x_47 = lean_ctor_get(x_19, 1); -x_48 = lean_ctor_get(x_19, 2); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_19); -lean_inc(x_1); -lean_inc(x_46); -x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_46, x_1); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_50 = lean_unsigned_to_nat(1u); -x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_46, x_1, x_50); -x_52 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_47); -lean_ctor_set(x_52, 2, x_48); -lean_ctor_set(x_18, 4, x_52); -x_53 = lean_st_ref_set(x_3, x_18, x_20); -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; -x_58 = lean_ctor_get(x_49, 0); -lean_inc(x_58); -lean_dec(x_49); -x_59 = lean_unsigned_to_nat(1u); -x_60 = lean_nat_add(x_58, x_59); -lean_dec(x_58); -x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_46, x_1, x_60); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_61); -lean_ctor_set(x_62, 1, x_47); -lean_ctor_set(x_62, 2, x_48); -lean_ctor_set(x_18, 4, x_62); -x_63 = lean_st_ref_set(x_3, x_18, x_20); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_65 = x_63; -} else { - lean_dec_ref(x_63); - x_65 = lean_box(0); -} -x_66 = lean_box(0); -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_65; -} -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_64); -return x_67; -} -} -} -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; -x_68 = lean_ctor_get(x_18, 0); -x_69 = lean_ctor_get(x_18, 1); -x_70 = lean_ctor_get(x_18, 2); -x_71 = lean_ctor_get(x_18, 3); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_18); -x_72 = lean_ctor_get(x_19, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_19, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_19, 2); -lean_inc(x_74); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - lean_ctor_release(x_19, 2); - x_75 = x_19; -} else { - lean_dec_ref(x_19); - x_75 = lean_box(0); -} -lean_inc(x_1); -lean_inc(x_72); -x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_72, x_1); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_77 = lean_unsigned_to_nat(1u); -x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_72, x_1, x_77); -if (lean_is_scalar(x_75)) { - x_79 = lean_alloc_ctor(0, 3, 0); -} else { - x_79 = x_75; -} -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_73); -lean_ctor_set(x_79, 2, x_74); -x_80 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_80, 0, x_68); -lean_ctor_set(x_80, 1, x_69); -lean_ctor_set(x_80, 2, x_70); -lean_ctor_set(x_80, 3, x_71); -lean_ctor_set(x_80, 4, x_79); -x_81 = lean_st_ref_set(x_3, x_80, x_20); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_83 = x_81; -} else { - lean_dec_ref(x_81); - x_83 = lean_box(0); -} -x_84 = lean_box(0); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_83; -} -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; 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_76, 0); -lean_inc(x_86); -lean_dec(x_76); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_add(x_86, x_87); -lean_dec(x_86); -x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_72, x_1, x_88); -if (lean_is_scalar(x_75)) { - x_90 = lean_alloc_ctor(0, 3, 0); -} else { - x_90 = x_75; -} -lean_ctor_set(x_90, 0, x_89); -lean_ctor_set(x_90, 1, x_73); -lean_ctor_set(x_90, 2, x_74); -x_91 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_91, 0, x_68); -lean_ctor_set(x_91, 1, x_69); -lean_ctor_set(x_91, 2, x_70); -lean_ctor_set(x_91, 3, x_71); -lean_ctor_set(x_91, 4, x_90); -x_92 = lean_st_ref_set(x_3, x_91, x_20); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); -} -x_95 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_94; -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; -x_6 = l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_recordUnfold___spec__3(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -size_t x_4; lean_object* x_5; -x_4 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_5 = l_Lean_PersistentHashMap_findAux___at_Lean_Meta_recordUnfold___spec__2(x_1, x_4, x_3); -lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___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: -{ -size_t x_7; lean_object* x_8; -x_7 = lean_unbox_usize(x_1); -lean_dec(x_1); -x_8 = l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_recordUnfold___spec__6(x_7, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_7 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_8 = l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_recordUnfold___spec__5(x_1, x_6, x_7, x_4, x_5); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_recordUnfold(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_Lean_Meta_recordDefEqHeuristic(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_1); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_7, 0); -lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_16 = lean_ctor_get(x_7, 1); -lean_inc(x_16); -lean_dec(x_7); -x_17 = lean_st_ref_take(x_3, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 4); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_18, 4); -lean_dec(x_22); -x_23 = !lean_is_exclusive(x_19); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_19, 1); -lean_inc(x_1); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_24, x_1); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_26); -lean_ctor_set(x_19, 1, x_27); -x_28 = lean_st_ref_set(x_3, x_18, x_20); -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; uint8_t x_40; -x_35 = lean_ctor_get(x_25, 0); -lean_inc(x_35); -lean_dec(x_25); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_35, x_36); -lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_37); -lean_ctor_set(x_19, 1, x_38); -x_39 = lean_st_ref_set(x_3, x_18, x_20); -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; lean_object* x_49; -x_46 = lean_ctor_get(x_19, 0); -x_47 = lean_ctor_get(x_19, 1); -x_48 = lean_ctor_get(x_19, 2); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_19); -lean_inc(x_1); -lean_inc(x_47); -x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_47, x_1); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_50 = lean_unsigned_to_nat(1u); -x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_47, x_1, x_50); -x_52 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_52, 0, x_46); -lean_ctor_set(x_52, 1, x_51); -lean_ctor_set(x_52, 2, x_48); -lean_ctor_set(x_18, 4, x_52); -x_53 = lean_st_ref_set(x_3, x_18, x_20); -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; -x_58 = lean_ctor_get(x_49, 0); -lean_inc(x_58); -lean_dec(x_49); -x_59 = lean_unsigned_to_nat(1u); -x_60 = lean_nat_add(x_58, x_59); -lean_dec(x_58); -x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_47, x_1, x_60); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_46); -lean_ctor_set(x_62, 1, x_61); -lean_ctor_set(x_62, 2, x_48); -lean_ctor_set(x_18, 4, x_62); -x_63 = lean_st_ref_set(x_3, x_18, x_20); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_65 = x_63; -} else { - lean_dec_ref(x_63); - x_65 = lean_box(0); -} -x_66 = lean_box(0); -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_65; -} -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_64); -return x_67; -} -} -} -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; -x_68 = lean_ctor_get(x_18, 0); -x_69 = lean_ctor_get(x_18, 1); -x_70 = lean_ctor_get(x_18, 2); -x_71 = lean_ctor_get(x_18, 3); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_18); -x_72 = lean_ctor_get(x_19, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_19, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_19, 2); -lean_inc(x_74); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - lean_ctor_release(x_19, 2); - x_75 = x_19; -} else { - lean_dec_ref(x_19); - x_75 = lean_box(0); -} -lean_inc(x_1); -lean_inc(x_73); -x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_73, x_1); -if (lean_obj_tag(x_76) == 0) -{ -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_77 = lean_unsigned_to_nat(1u); -x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_73, x_1, x_77); -if (lean_is_scalar(x_75)) { - x_79 = lean_alloc_ctor(0, 3, 0); -} else { - x_79 = x_75; -} -lean_ctor_set(x_79, 0, x_72); -lean_ctor_set(x_79, 1, x_78); -lean_ctor_set(x_79, 2, x_74); -x_80 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_80, 0, x_68); -lean_ctor_set(x_80, 1, x_69); -lean_ctor_set(x_80, 2, x_70); -lean_ctor_set(x_80, 3, x_71); -lean_ctor_set(x_80, 4, x_79); -x_81 = lean_st_ref_set(x_3, x_80, x_20); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_83 = x_81; -} else { - lean_dec_ref(x_81); - x_83 = lean_box(0); -} -x_84 = lean_box(0); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_83; -} -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; -} -else -{ -lean_object* x_86; lean_object* x_87; 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_76, 0); -lean_inc(x_86); -lean_dec(x_76); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_add(x_86, x_87); -lean_dec(x_86); -x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_73, x_1, x_88); -if (lean_is_scalar(x_75)) { - x_90 = lean_alloc_ctor(0, 3, 0); -} else { - x_90 = x_75; -} -lean_ctor_set(x_90, 0, x_72); -lean_ctor_set(x_90, 1, x_89); -lean_ctor_set(x_90, 2, x_74); -x_91 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_91, 0, x_68); -lean_ctor_set(x_91, 1, x_69); -lean_ctor_set(x_91, 2, x_70); -lean_ctor_set(x_91, 3, x_71); -lean_ctor_set(x_91, 4, x_90); -x_92 = lean_st_ref_set(x_3, x_91, x_20); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); -} -x_95 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_94; -} -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_recordDefEqHeuristic___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l_Lean_Meta_recordDefEqHeuristic(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_Lean_Meta_recordInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_unbox(x_8); -lean_dec(x_8); -if (x_9 == 0) -{ -uint8_t x_10; -lean_dec(x_1); -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_7, 0); -lean_dec(x_11); -x_12 = lean_box(0); -lean_ctor_set(x_7, 0, x_12); -return x_7; -} -else -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_7, 1); -lean_inc(x_13); -lean_dec(x_7); -x_14 = lean_box(0); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_14); -lean_ctor_set(x_15, 1, x_13); -return x_15; -} -} -else -{ -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_16 = lean_ctor_get(x_7, 1); -lean_inc(x_16); -lean_dec(x_7); -x_17 = lean_st_ref_take(x_3, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_18, 4); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 1); -lean_inc(x_20); -lean_dec(x_17); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) -{ -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_18, 4); -lean_dec(x_22); -x_23 = !lean_is_exclusive(x_19); -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; -x_24 = lean_ctor_get(x_19, 2); -lean_inc(x_1); -lean_inc(x_24); -x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_24, x_1); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_unsigned_to_nat(1u); -x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_26); -lean_ctor_set(x_19, 2, x_27); -x_28 = lean_st_ref_set(x_3, x_18, x_20); -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; uint8_t x_40; -x_35 = lean_ctor_get(x_25, 0); -lean_inc(x_35); -lean_dec(x_25); -x_36 = lean_unsigned_to_nat(1u); -x_37 = lean_nat_add(x_35, x_36); -lean_dec(x_35); -x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_24, x_1, x_37); -lean_ctor_set(x_19, 2, x_38); -x_39 = lean_st_ref_set(x_3, x_18, x_20); -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; lean_object* x_49; -x_46 = lean_ctor_get(x_19, 0); -x_47 = lean_ctor_get(x_19, 1); -x_48 = lean_ctor_get(x_19, 2); -lean_inc(x_48); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_19); -lean_inc(x_1); -lean_inc(x_48); -x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_48, x_1); -if (lean_obj_tag(x_49) == 0) -{ -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_50 = lean_unsigned_to_nat(1u); -x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_48, x_1, x_50); -x_52 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_52, 0, x_46); -lean_ctor_set(x_52, 1, x_47); -lean_ctor_set(x_52, 2, x_51); -lean_ctor_set(x_18, 4, x_52); -x_53 = lean_st_ref_set(x_3, x_18, x_20); -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_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_54); -return x_57; } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__12() { +_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_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_58 = lean_ctor_get(x_49, 0); -lean_inc(x_58); -lean_dec(x_49); -x_59 = lean_unsigned_to_nat(1u); -x_60 = lean_nat_add(x_58, x_59); -lean_dec(x_58); -x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_48, x_1, x_60); -x_62 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_62, 0, x_46); -lean_ctor_set(x_62, 1, x_47); -lean_ctor_set(x_62, 2, x_61); -lean_ctor_set(x_18, 4, x_62); -x_63 = lean_st_ref_set(x_3, x_18, x_20); -x_64 = lean_ctor_get(x_63, 1); -lean_inc(x_64); -if (lean_is_exclusive(x_63)) { - lean_ctor_release(x_63, 0); - lean_ctor_release(x_63, 1); - x_65 = x_63; -} else { - lean_dec_ref(x_63); - x_65 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Basic", 5); +return x_1; } -x_66 = lean_box(0); -if (lean_is_scalar(x_65)) { - x_67 = lean_alloc_ctor(0, 2, 0); -} else { - x_67 = x_65; } -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_64); -return x_67; +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__11; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__12; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__14() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_hyg", 4); +return x_1; +} } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__15() { +_start: { -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_68 = lean_ctor_get(x_18, 0); -x_69 = lean_ctor_get(x_18, 1); -x_70 = lean_ctor_get(x_18, 2); -x_71 = lean_ctor_get(x_18, 3); -lean_inc(x_71); -lean_inc(x_70); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_18); -x_72 = lean_ctor_get(x_19, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_19, 1); -lean_inc(x_73); -x_74 = lean_ctor_get(x_19, 2); -lean_inc(x_74); -if (lean_is_exclusive(x_19)) { - lean_ctor_release(x_19, 0); - lean_ctor_release(x_19, 1); - lean_ctor_release(x_19, 2); - x_75 = x_19; -} else { - lean_dec_ref(x_19); - x_75 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__13; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__14; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } -lean_inc(x_1); -lean_inc(x_74); -x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_recordUnfold___spec__1(x_74, x_1); -if (lean_obj_tag(x_76) == 0) +} +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__16() { +_start: { -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_77 = lean_unsigned_to_nat(1u); -x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_74, x_1, x_77); -if (lean_is_scalar(x_75)) { - x_79 = lean_alloc_ctor(0, 3, 0); -} else { - x_79 = x_75; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__15; +x_2 = lean_unsigned_to_nat(1758u); +x_3 = l_Lean_Name_num___override(x_1, x_2); +return x_3; } -lean_ctor_set(x_79, 0, x_72); -lean_ctor_set(x_79, 1, x_73); -lean_ctor_set(x_79, 2, x_78); -x_80 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_80, 0, x_68); -lean_ctor_set(x_80, 1, x_69); -lean_ctor_set(x_80, 2, x_70); -lean_ctor_set(x_80, 3, x_71); -lean_ctor_set(x_80, 4, x_79); -x_81 = lean_st_ref_set(x_3, x_80, x_20); -x_82 = lean_ctor_get(x_81, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_83 = x_81; -} else { - lean_dec_ref(x_81); - x_83 = lean_box(0); } -x_84 = lean_box(0); -if (lean_is_scalar(x_83)) { - x_85 = lean_alloc_ctor(0, 2, 0); -} else { - x_85 = x_83; +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__17() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("debug", 5); +return x_1; } -lean_ctor_set(x_85, 0, x_84); -lean_ctor_set(x_85, 1, x_82); -return x_85; } -else +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__18() { +_start: { -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_76, 0); -lean_inc(x_86); -lean_dec(x_76); -x_87 = lean_unsigned_to_nat(1u); -x_88 = lean_nat_add(x_86, x_87); -lean_dec(x_86); -x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Meta_recordUnfold___spec__4(x_74, x_1, x_88); -if (lean_is_scalar(x_75)) { - x_90 = lean_alloc_ctor(0, 3, 0); -} else { - x_90 = x_75; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__17; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -lean_ctor_set(x_90, 0, x_72); -lean_ctor_set(x_90, 1, x_73); -lean_ctor_set(x_90, 2, x_89); -x_91 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_91, 0, x_68); -lean_ctor_set(x_91, 1, x_69); -lean_ctor_set(x_91, 2, x_70); -lean_ctor_set(x_91, 3, x_71); -lean_ctor_set(x_91, 4, x_90); -x_92 = lean_st_ref_set(x_3, x_91, x_20); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_94 = x_92; -} else { - lean_dec_ref(x_92); - x_94 = lean_box(0); } -x_95 = lean_box(0); -if (lean_is_scalar(x_94)) { - x_96 = lean_alloc_ctor(0, 2, 0); -} else { - x_96 = x_94; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758_(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__2; +x_3 = 0; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__16; +x_5 = l_Lean_registerTraceClass(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; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_Meta_initFn____x40_Lean_Meta_Basic___hyg_1758____closed__18; +x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); +return x_8; } -lean_ctor_set(x_96, 0, x_95); -lean_ctor_set(x_96, 1, x_93); -return x_96; +else +{ +uint8_t x_9; +x_9 = !lean_is_exclusive(x_5); +if (x_9 == 0) +{ +return x_5; } +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_ctor_get(x_5, 0); +x_11 = lean_ctor_get(x_5, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_5); +x_12 = lean_alloc_ctor(1, 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_Meta_recordInstance___boxed(lean_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_liftMetaM___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_7; -x_7 = l_Lean_Meta_recordInstance(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_3; +x_3 = lean_apply_2(x_1, lean_box(0), x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___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_Lean_Meta_liftMetaM(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_eq(x_3, x_4); -if (x_6 == 0) +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_liftMetaM___rarg), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_7; -x_7 = lean_array_uget(x_2, x_3); -switch (lean_obj_tag(x_7)) { -case 0: +lean_object* x_9; lean_object* x_10; +x_9 = lean_apply_2(x_3, lean_box(0), x_1); +x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM___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_8; lean_object* x_9; lean_object* x_10; -x_8 = lean_ctor_get(x_7, 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; +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +lean_dec(x_2); +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_mapMetaM___rarg___lambda__1), 8, 2); +lean_closure_set(x_7, 0, x_5); +lean_closure_set(x_7, 1, x_3); +x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -lean_inc(x_1); -x_10 = lean_apply_3(x_1, x_5, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -uint8_t x_11; +x_9 = lean_apply_2(x_8, lean_box(0), x_7); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); lean_dec(x_1); -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -return x_10; +x_11 = lean_apply_1(x_10, lean_box(0)); +x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_mapMetaM(lean_object* x_1) { +_start: { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_13, 0, x_12); -return x_13; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_mapMetaM___rarg), 5, 0); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_14; size_t x_15; size_t x_16; -x_14 = lean_ctor_get(x_10, 0); -lean_inc(x_14); -lean_dec(x_10); -x_15 = 1; -x_16 = lean_usize_add(x_3, x_15); -x_3 = x_16; -x_5 = x_14; -goto _start; +lean_object* x_9; lean_object* x_10; +x_9 = lean_apply_1(x_1, x_3); +x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -case 1: +LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___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_18; lean_object* x_19; -x_18 = lean_ctor_get(x_7, 0); -lean_inc(x_18); -lean_dec(x_7); -lean_inc(x_1); -x_19 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(x_1, x_18, x_5); -if (lean_obj_tag(x_19) == 0) +lean_object* x_9; lean_object* x_10; +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg___lambda__1), 8, 2); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_3); +x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: { -uint8_t x_20; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +lean_dec(x_2); +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg___lambda__2), 8, 2); +lean_closure_set(x_7, 0, x_5); +lean_closure_set(x_7, 1, x_3); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_apply_2(x_8, lean_box(0), x_7); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); lean_dec(x_1); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -return x_19; +x_11 = lean_apply_1(x_10, lean_box(0)); +x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_Meta_map1MetaM(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -return x_22; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Meta_map1MetaM___rarg), 5, 0); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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) { +_start: { -lean_object* x_23; size_t x_24; size_t x_25; -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -lean_dec(x_19); -x_24 = 1; -x_25 = lean_usize_add(x_3, x_24); -x_3 = x_25; -x_5 = x_23; -goto _start; +lean_object* x_10; lean_object* x_11; +x_10 = lean_apply_2(x_1, x_3, x_4); +x_11 = lean_apply_7(x_2, lean_box(0), x_10, x_5, x_6, x_7, x_8, x_9); +return x_11; } } -default: +LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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: { -size_t x_27; size_t x_28; -x_27 = 1; -x_28 = lean_usize_add(x_3, x_27); -x_3 = x_28; -goto _start; -} +lean_object* x_9; lean_object* x_10; +x_9 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg___lambda__1), 9, 2); +lean_closure_set(x_9, 0, x_1); +lean_closure_set(x_9, 1, x_3); +x_10 = lean_apply_7(x_2, lean_box(0), x_9, x_4, x_5, x_6, x_7, x_8); +return x_10; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM___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_30; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_2, 1); +lean_inc(x_6); +lean_dec(x_2); +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg___lambda__2), 8, 2); +lean_closure_set(x_7, 0, x_5); +lean_closure_set(x_7, 1, x_3); +x_8 = lean_ctor_get(x_1, 0); +lean_inc(x_8); +x_9 = lean_apply_2(x_8, lean_box(0), x_7); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); lean_dec(x_1); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_5); -return x_30; -} +x_11 = lean_apply_1(x_10, lean_box(0)); +x_12 = lean_apply_4(x_6, lean_box(0), lean_box(0), x_9, x_11); +return x_12; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_map2MetaM(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed), 5, 0); +x_4 = lean_alloc_closure((void*)(l_Lean_Meta_map2MetaM___rarg), 5, 0); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_7; uint8_t x_8; -x_7 = lean_array_get_size(x_2); -x_8 = lean_nat_dec_lt(x_5, x_7); +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_7 = lean_st_ref_take(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); lean_dec(x_7); -if (x_8 == 0) -{ -lean_object* x_9; -lean_dec(x_5); -lean_dec(x_1); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_6); -return x_9; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_array_fget(x_2, x_5); -x_11 = lean_array_fget(x_3, x_5); -lean_inc(x_1); -x_12 = lean_apply_3(x_1, x_6, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -uint8_t x_13; -lean_dec(x_5); -lean_dec(x_1); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 1); +x_12 = lean_apply_1(x_1, x_11); +lean_ctor_set(x_8, 1, x_12); +x_13 = lean_st_ref_set(x_3, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -return x_12; +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +x_16 = lean_box(0); +lean_ctor_set(x_13, 0, x_16); +return x_13; } else { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_15, 0, x_14); -return x_15; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_13, 1); +lean_inc(x_17); +lean_dec(x_13); +x_18 = lean_box(0); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +return x_19; } } else { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 0); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_add(x_5, x_17); -lean_dec(x_5); -x_4 = lean_box(0); -x_5 = x_18; -x_6 = x_16; -goto _start; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_20 = lean_ctor_get(x_8, 0); +x_21 = lean_ctor_get(x_8, 1); +x_22 = lean_ctor_get(x_8, 2); +x_23 = lean_ctor_get(x_8, 3); +x_24 = lean_ctor_get(x_8, 4); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_dec(x_8); +x_25 = lean_apply_1(x_1, x_21); +x_26 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_22); +lean_ctor_set(x_26, 3, x_23); +lean_ctor_set(x_26, 4, x_24); +x_27 = lean_st_ref_set(x_3, x_26, x_9); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_29 = x_27; +} else { + lean_dec_ref(x_27); + x_29 = lean_box(0); +} +x_30 = lean_box(0); +if (lean_is_scalar(x_29)) { + x_31 = lean_alloc_ctor(0, 2, 0); +} else { + x_31 = x_29; } +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_28); +return x_31; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyCache___boxed(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_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed), 6, 0); -return x_4; +lean_object* x_7; +x_7 = l_Lean_Meta_modifyCache(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_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache(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_2) == 0) -{ -lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_4 = lean_ctor_get(x_2, 0); -lean_inc(x_4); -lean_dec(x_2); -x_5 = lean_array_get_size(x_4); -x_6 = lean_unsigned_to_nat(0u); -x_7 = lean_nat_dec_lt(x_6, x_5); -if (x_7 == 0) +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_st_ref_take(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_8; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_8 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_8, 0, x_3); -return x_8; -} -else +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_8, 1); +lean_dec(x_12); +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) { -uint8_t x_9; -x_9 = lean_nat_dec_le(x_5, x_5); -if (x_9 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_9, 0); +x_15 = lean_apply_1(x_1, x_14); +lean_ctor_set(x_9, 0, x_15); +x_16 = lean_st_ref_set(x_3, x_8, x_10); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_10; -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_3); -return x_10; +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 { -size_t x_11; size_t x_12; lean_object* x_13; -x_11 = 0; -x_12 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(x_1, x_4, x_11, x_12, x_3); -lean_dec(x_4); -return x_13; -} +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 { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_ctor_get(x_2, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_2, 1); -lean_inc(x_15); -lean_dec(x_2); -x_16 = lean_unsigned_to_nat(0u); -x_17 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(x_1, x_14, x_15, lean_box(0), x_16, x_3); -lean_dec(x_15); -lean_dec(x_14); -return x_17; -} +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_ctor_get(x_9, 3); +x_27 = lean_ctor_get(x_9, 4); +x_28 = lean_ctor_get(x_9, 5); +x_29 = lean_ctor_get(x_9, 6); +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_dec(x_9); +x_30 = lean_apply_1(x_1, x_23); +x_31 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_24); +lean_ctor_set(x_31, 2, x_25); +lean_ctor_set(x_31, 3, x_26); +lean_ctor_set(x_31, 4, x_27); +lean_ctor_set(x_31, 5, x_28); +lean_ctor_set(x_31, 6, x_29); +lean_ctor_set(x_8, 1, x_31); +x_32 = lean_st_ref_set(x_3, x_8, x_10); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); } +x_35 = lean_box(0); +if (lean_is_scalar(x_34)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_34; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg), 3, 0); -return x_4; +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_1, 0); -lean_inc(x_4); -lean_dec(x_1); -x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(x_2, x_4, x_3); -return x_5; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); +x_38 = lean_ctor_get(x_8, 2); +x_39 = lean_ctor_get(x_8, 3); +x_40 = lean_ctor_get(x_8, 4); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_8); +x_41 = lean_ctor_get(x_9, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_9, 1); +lean_inc(x_42); +x_43 = lean_ctor_get(x_9, 2); +lean_inc(x_43); +x_44 = lean_ctor_get(x_9, 3); +lean_inc(x_44); +x_45 = lean_ctor_get(x_9, 4); +lean_inc(x_45); +x_46 = lean_ctor_get(x_9, 5); +lean_inc(x_46); +x_47 = lean_ctor_get(x_9, 6); +lean_inc(x_47); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + lean_ctor_release(x_9, 2); + lean_ctor_release(x_9, 3); + lean_ctor_release(x_9, 4); + lean_ctor_release(x_9, 5); + lean_ctor_release(x_9, 6); + x_48 = x_9; +} else { + lean_dec_ref(x_9); + x_48 = lean_box(0); } +x_49 = lean_apply_1(x_1, x_41); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(0, 7, 0); +} else { + x_50 = x_48; } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_5, 0, x_3); -lean_ctor_set(x_5, 1, x_4); -x_6 = lean_apply_2(x_1, x_5, x_2); -if (lean_obj_tag(x_6) == 0) -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_8, 0, x_7); -return x_8; +lean_ctor_set(x_50, 0, x_49); +lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 2, x_43); +lean_ctor_set(x_50, 3, x_44); +lean_ctor_set(x_50, 4, x_45); +lean_ctor_set(x_50, 5, x_46); +lean_ctor_set(x_50, 6, x_47); +x_51 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_51, 0, x_37); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_38); +lean_ctor_set(x_51, 3, x_39); +lean_ctor_set(x_51, 4, x_40); +x_52 = lean_st_ref_set(x_3, x_51, x_10); +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); } -else -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_6, 0); -lean_inc(x_9); -lean_dec(x_6); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -return x_10; +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; } } } -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyInferTypeCache___boxed(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_4; lean_object* x_5; lean_object* x_6; -x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___lambda__1), 4, 1); -lean_closure_set(x_4, 0, x_3); -x_5 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(x_1, x_4, x_2); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); +lean_object* x_7; +x_7 = l_Lean_Meta_modifyInferTypeCache(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); -return x_6; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_7; } } -LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache(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; lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_3 = lean_ctor_get(x_1, 0); -x_4 = lean_ctor_get(x_1, 1); -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -x_7 = lean_nat_dec_eq(x_4, x_6); -if (x_7 == 0) -{ -uint8_t x_8; -x_8 = lean_nat_dec_lt(x_6, x_4); -return x_8; -} -else -{ -uint8_t x_9; -x_9 = l_Lean_Name_lt(x_3, x_5); -return x_9; -} -} -} -static lean_object* _init_l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___closed__1() { -_start: +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_st_ref_take(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1___boxed), 2, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_8, 1); +lean_dec(x_12); +x_13 = !lean_is_exclusive(x_9); +if (x_13 == 0) { -uint8_t x_4; -x_4 = lean_nat_dec_lt(x_2, x_3); -if (x_4 == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_14 = lean_ctor_get(x_9, 5); +x_15 = lean_apply_1(x_1, x_14); +lean_ctor_set(x_9, 5, x_15); +x_16 = lean_st_ref_set(x_3, x_8, x_10); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_dec(x_2); -return x_1; +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_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___closed__1; -lean_inc(x_2); -x_6 = l_Array_qpartition___rarg(x_1, x_5, x_2, x_3); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = lean_nat_dec_le(x_3, x_7); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(x_8, x_2, x_7); -x_11 = lean_unsigned_to_nat(1u); -x_12 = lean_nat_add(x_7, x_11); -lean_dec(x_7); -x_1 = x_10; -x_2 = x_12; -goto _start; +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 { -lean_dec(x_7); -lean_dec(x_2); -return x_8; -} +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_ctor_get(x_9, 3); +x_27 = lean_ctor_get(x_9, 4); +x_28 = lean_ctor_get(x_9, 5); +x_29 = lean_ctor_get(x_9, 6); +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_dec(x_9); +x_30 = lean_apply_1(x_1, x_28); +x_31 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_31, 0, x_23); +lean_ctor_set(x_31, 1, x_24); +lean_ctor_set(x_31, 2, x_25); +lean_ctor_set(x_31, 3, x_26); +lean_ctor_set(x_31, 4, x_27); +lean_ctor_set(x_31, 5, x_30); +lean_ctor_set(x_31, 6, x_29); +lean_ctor_set(x_8, 1, x_31); +x_32 = lean_st_ref_set(x_3, x_8, x_10); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); } +x_35 = lean_box(0); +if (lean_is_scalar(x_34)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_34; } +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; } -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -uint8_t x_4; -x_4 = !lean_is_exclusive(x_2); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 0); -x_6 = lean_ctor_get(x_2, 1); -x_7 = lean_nat_dec_lt(x_1, x_6); -if (x_7 == 0) -{ -lean_object* x_8; -lean_free_object(x_2); -lean_dec(x_6); -lean_dec(x_5); -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; -x_9 = lean_array_push(x_3, x_2); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -return x_10; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); +x_38 = lean_ctor_get(x_8, 2); +x_39 = lean_ctor_get(x_8, 3); +x_40 = lean_ctor_get(x_8, 4); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_8); +x_41 = lean_ctor_get(x_9, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_9, 1); +lean_inc(x_42); +x_43 = lean_ctor_get(x_9, 2); +lean_inc(x_43); +x_44 = lean_ctor_get(x_9, 3); +lean_inc(x_44); +x_45 = lean_ctor_get(x_9, 4); +lean_inc(x_45); +x_46 = lean_ctor_get(x_9, 5); +lean_inc(x_46); +x_47 = lean_ctor_get(x_9, 6); +lean_inc(x_47); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + lean_ctor_release(x_9, 2); + lean_ctor_release(x_9, 3); + lean_ctor_release(x_9, 4); + lean_ctor_release(x_9, 5); + lean_ctor_release(x_9, 6); + x_48 = x_9; +} else { + lean_dec_ref(x_9); + x_48 = lean_box(0); } +x_49 = lean_apply_1(x_1, x_46); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(0, 7, 0); +} else { + x_50 = x_48; } -else -{ -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_2, 0); -x_12 = lean_ctor_get(x_2, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_2); -x_13 = lean_nat_dec_lt(x_1, x_12); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_12); -lean_dec(x_11); -x_14 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_14, 0, x_3); -return x_14; +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 2, x_43); +lean_ctor_set(x_50, 3, x_44); +lean_ctor_set(x_50, 4, x_45); +lean_ctor_set(x_50, 5, x_49); +lean_ctor_set(x_50, 6, x_47); +x_51 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_51, 0, x_37); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_38); +lean_ctor_set(x_51, 3, x_39); +lean_ctor_set(x_51, 4, x_40); +x_52 = lean_st_ref_set(x_3, x_51, x_10); +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); } -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_11); -lean_ctor_set(x_15, 1, x_12); -x_16 = lean_array_push(x_3, x_15); -x_17 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_17, 0, x_16); -return x_17; +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; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqTransientCache___boxed(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; lean_object* x_4; 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_alloc_closure((void*)(l_Lean_Meta_collectAboveThreshold___lambda__1___boxed), 3, 1); -lean_closure_set(x_3, 0, x_2); -x_4 = l_Lean_Meta_ParamInfo_backDeps___default___closed__1; -x_5 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1(x_1, x_4, x_3); -x_6 = lean_array_get_size(x_5); -x_7 = lean_unsigned_to_nat(1u); -x_8 = lean_nat_sub(x_6, x_7); -lean_dec(x_6); -x_9 = lean_unsigned_to_nat(0u); -x_10 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(x_5, x_9, x_8); -lean_dec(x_8); -return x_10; +lean_object* x_7; +x_7 = l_Lean_Meta_modifyDefEqTransientCache(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_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); -lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); -lean_dec(x_4); -x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); -lean_dec(x_2); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_st_ref_take(x_3, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_8, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_7, 1); +lean_inc(x_10); +lean_dec(x_7); +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_8, 1); +lean_dec(x_12); +x_13 = !lean_is_exclusive(x_9); +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_9, 6); +x_15 = lean_apply_1(x_1, x_14); +lean_ctor_set(x_9, 6, x_15); +x_16 = lean_st_ref_set(x_3, x_8, x_10); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_7; -x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_3); -lean_dec(x_2); -return x_7; -} +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; } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { -_start: +else { -uint8_t x_3; lean_object* x_4; -x_3 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___lambda__1(x_1, x_2); -lean_dec(x_2); -lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +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; } } -LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +else { -lean_object* x_4; -x_4 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(x_1, x_2, x_3); -lean_dec(x_3); -return x_4; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; 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_23 = lean_ctor_get(x_9, 0); +x_24 = lean_ctor_get(x_9, 1); +x_25 = lean_ctor_get(x_9, 2); +x_26 = lean_ctor_get(x_9, 3); +x_27 = lean_ctor_get(x_9, 4); +x_28 = lean_ctor_get(x_9, 5); +x_29 = lean_ctor_get(x_9, 6); +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_dec(x_9); +x_30 = lean_apply_1(x_1, x_29); +x_31 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_31, 0, x_23); +lean_ctor_set(x_31, 1, x_24); +lean_ctor_set(x_31, 2, x_25); +lean_ctor_set(x_31, 3, x_26); +lean_ctor_set(x_31, 4, x_27); +lean_ctor_set(x_31, 5, x_28); +lean_ctor_set(x_31, 6, x_30); +lean_ctor_set(x_8, 1, x_31); +x_32 = lean_st_ref_set(x_3, x_8, x_10); +x_33 = lean_ctor_get(x_32, 1); +lean_inc(x_33); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_34 = x_32; +} else { + lean_dec_ref(x_32); + x_34 = lean_box(0); } +x_35 = lean_box(0); +if (lean_is_scalar(x_34)) { + x_36 = lean_alloc_ctor(0, 2, 0); +} else { + x_36 = x_34; } -LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Meta_collectAboveThreshold___lambda__1(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_33); +return x_36; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\n", 1); -return x_1; -} +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_37 = lean_ctor_get(x_8, 0); +x_38 = lean_ctor_get(x_8, 2); +x_39 = lean_ctor_get(x_8, 3); +x_40 = lean_ctor_get(x_8, 4); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_8); +x_41 = lean_ctor_get(x_9, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_9, 1); +lean_inc(x_42); +x_43 = lean_ctor_get(x_9, 2); +lean_inc(x_43); +x_44 = lean_ctor_get(x_9, 3); +lean_inc(x_44); +x_45 = lean_ctor_get(x_9, 4); +lean_inc(x_45); +x_46 = lean_ctor_get(x_9, 5); +lean_inc(x_46); +x_47 = lean_ctor_get(x_9, 6); +lean_inc(x_47); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + lean_ctor_release(x_9, 2); + lean_ctor_release(x_9, 3); + lean_ctor_release(x_9, 4); + lean_ctor_release(x_9, 5); + lean_ctor_release(x_9, 6); + x_48 = x_9; +} else { + lean_dec_ref(x_9); + x_48 = lean_box(0); } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +x_49 = lean_apply_1(x_1, x_47); +if (lean_is_scalar(x_48)) { + x_50 = lean_alloc_ctor(0, 7, 0); +} else { + x_50 = x_48; } +lean_ctor_set(x_50, 0, x_41); +lean_ctor_set(x_50, 1, x_42); +lean_ctor_set(x_50, 2, x_43); +lean_ctor_set(x_50, 3, x_44); +lean_ctor_set(x_50, 4, x_45); +lean_ctor_set(x_50, 5, x_46); +lean_ctor_set(x_50, 6, x_49); +x_51 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_51, 0, x_37); +lean_ctor_set(x_51, 1, x_50); +lean_ctor_set(x_51, 2, x_38); +lean_ctor_set(x_51, 3, x_39); +lean_ctor_set(x_51, 4, x_40); +x_52 = lean_st_ref_set(x_3, x_51, x_10); +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); } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes(" ↦ ", 5); -return x_1; +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; } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5() { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDefEqPermCache___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_7; +x_7 = l_Lean_Meta_modifyDefEqPermCache(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_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___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_11; -x_11 = lean_usize_dec_lt(x_3, x_2); -if (x_11 == 0) +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_st_ref_take(x_1, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +x_8 = lean_ctor_get(x_5, 1); +lean_inc(x_8); +lean_dec(x_5); +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) { -return x_4; -} -else +lean_object* x_10; uint8_t x_11; +x_10 = lean_ctor_get(x_6, 1); +lean_dec(x_10); +x_11 = !lean_is_exclusive(x_7); +if (x_11 == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_array_uget(x_1, x_3); -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_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_12 = lean_ctor_get(x_7, 6); lean_dec(x_12); -if (lean_obj_tag(x_4) == 0) -{ -lean_object* x_30; -x_30 = lean_ctor_get(x_4, 0); -lean_inc(x_30); -if (lean_obj_tag(x_30) == 0) -{ -uint8_t x_31; -x_31 = !lean_is_exclusive(x_4); -if (x_31 == 0) +x_13 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; +lean_ctor_set(x_7, 6, x_13); +x_14 = lean_st_ref_set(x_1, x_6, x_8); +x_15 = !lean_is_exclusive(x_14); +if (x_15 == 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; -x_32 = lean_ctor_get(x_4, 0); -lean_dec(x_32); -x_33 = l_Lean_MessageData_ofName(x_13); -x_34 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; -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_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__4; -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_Init_Data_Repr_0__Nat_reprFast(x_14); -x_39 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_4, 0, x_39); -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_4); -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_34); -x_42 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_42, 0, x_41); -x_5 = x_42; -goto block_10; +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; } 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_dec(x_4); -x_43 = l_Lean_MessageData_ofName(x_13); -x_44 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); -x_46 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__4; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); -x_49 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_49, 0, x_48); -x_50 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_50, 0, x_49); -x_51 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_51, 0, x_47); -lean_ctor_set(x_51, 1, x_50); -x_52 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_44); -x_53 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_53, 0, x_52); -x_5 = x_53; -goto block_10; +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; } } else { -lean_object* x_54; -lean_dec(x_30); -x_54 = lean_box(0); -x_15 = x_54; -goto block_29; +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; 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_ctor_get(x_7, 0); +x_22 = lean_ctor_get(x_7, 1); +x_23 = lean_ctor_get(x_7, 2); +x_24 = lean_ctor_get(x_7, 3); +x_25 = lean_ctor_get(x_7, 4); +x_26 = lean_ctor_get(x_7, 5); +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_dec(x_7); +x_27 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; +x_28 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_28, 0, x_21); +lean_ctor_set(x_28, 1, x_22); +lean_ctor_set(x_28, 2, x_23); +lean_ctor_set(x_28, 3, x_24); +lean_ctor_set(x_28, 4, x_25); +lean_ctor_set(x_28, 5, x_26); +lean_ctor_set(x_28, 6, x_27); +lean_ctor_set(x_6, 1, x_28); +x_29 = lean_st_ref_set(x_1, x_6, x_8); +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_31 = x_29; +} else { + lean_dec_ref(x_29); + x_31 = lean_box(0); +} +x_32 = lean_box(0); +if (lean_is_scalar(x_31)) { + x_33 = lean_alloc_ctor(0, 2, 0); +} else { + x_33 = x_31; +} +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_30); +return x_33; } } else { -lean_object* x_55; -x_55 = lean_box(0); -x_15 = x_55; -goto block_29; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_34 = lean_ctor_get(x_6, 0); +x_35 = lean_ctor_get(x_6, 2); +x_36 = lean_ctor_get(x_6, 3); +x_37 = lean_ctor_get(x_6, 4); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_6); +x_38 = lean_ctor_get(x_7, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_7, 1); +lean_inc(x_39); +x_40 = lean_ctor_get(x_7, 2); +lean_inc(x_40); +x_41 = lean_ctor_get(x_7, 3); +lean_inc(x_41); +x_42 = lean_ctor_get(x_7, 4); +lean_inc(x_42); +x_43 = lean_ctor_get(x_7, 5); +lean_inc(x_43); +if (lean_is_exclusive(x_7)) { + lean_ctor_release(x_7, 0); + lean_ctor_release(x_7, 1); + lean_ctor_release(x_7, 2); + lean_ctor_release(x_7, 3); + lean_ctor_release(x_7, 4); + lean_ctor_release(x_7, 5); + lean_ctor_release(x_7, 6); + x_44 = x_7; +} else { + lean_dec_ref(x_7); + x_44 = lean_box(0); } -block_29: -{ -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_dec(x_15); -x_16 = l_Lean_MessageData_ofName(x_13); -x_17 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__2; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__4; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -x_21 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); -x_22 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_23, 0, x_22); -x_24 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_24, 0, x_20); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_4); -lean_ctor_set(x_27, 1, x_26); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_27); -x_5 = x_28; -goto block_10; +x_45 = l_Lean_Meta_instInhabitedDefEqCache___closed__1; +if (lean_is_scalar(x_44)) { + x_46 = lean_alloc_ctor(0, 7, 0); +} else { + x_46 = x_44; } +lean_ctor_set(x_46, 0, x_38); +lean_ctor_set(x_46, 1, x_39); +lean_ctor_set(x_46, 2, x_40); +lean_ctor_set(x_46, 3, x_41); +lean_ctor_set(x_46, 4, x_42); +lean_ctor_set(x_46, 5, x_43); +lean_ctor_set(x_46, 6, x_45); +x_47 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_47, 0, x_34); +lean_ctor_set(x_47, 1, x_46); +lean_ctor_set(x_47, 2, x_35); +lean_ctor_set(x_47, 3, x_36); +lean_ctor_set(x_47, 4, x_37); +x_48 = lean_st_ref_set(x_1, x_47, x_8); +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); } -block_10: -{ -lean_object* x_6; size_t x_7; size_t x_8; -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -x_7 = 1; -x_8 = lean_usize_add(x_3, x_7); -x_3 = x_8; -x_4 = x_6; -goto _start; +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; } } -LEAN_EXPORT lean_object* l_Lean_Meta_mkMessageBodyFor_x3f(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; uint8_t x_4; -x_3 = l_Lean_Meta_collectAboveThreshold(x_1, x_2); -x_4 = l_Array_isEmpty___rarg(x_3); -if (x_4 == 0) -{ -lean_object* x_5; size_t x_6; size_t x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_5 = lean_array_get_size(x_3); -x_6 = lean_usize_of_nat(x_5); -lean_dec(x_5); -x_7 = 0; -x_8 = l_Lean_MessageData_nil; -x_9 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1(x_3, x_6, x_7, x_8); -lean_dec(x_3); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -return x_10; } -else +LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches(lean_object* x_1) { +_start: { -lean_object* x_11; -lean_dec(x_3); -x_11 = lean_box(0); -return x_11; -} +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_resetDefEqPermCaches___rarg___boxed), 4, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_5; size_t x_6; lean_object* x_7; -x_5 = lean_unbox_usize(x_2); -lean_dec(x_2); -x_6 = lean_unbox_usize(x_3); +lean_object* x_5; +x_5 = l_Lean_Meta_resetDefEqPermCaches___rarg(x_1, x_2, x_3, x_4); lean_dec(x_3); -x_7 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1(x_1, x_5, x_6, x_4); +lean_dec(x_2); lean_dec(x_1); -return x_7; +return x_5; } } -static lean_object* _init_l_Lean_Meta_appendOptMessageData___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_resetDefEqPermCaches___boxed(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); +lean_object* x_2; +x_2 = l_Lean_Meta_resetDefEqPermCaches(x_1); +lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_appendOptMessageData___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDiag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_appendOptMessageData___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_Meta_appendOptMessageData(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -if (lean_obj_tag(x_3) == 0) +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_dec(x_2); -return x_1; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_4; lean_object* x_5; -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -lean_dec(x_3); -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_14; -x_14 = lean_ctor_get(x_1, 0); -lean_inc(x_14); -if (lean_obj_tag(x_14) == 0) -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_1); -if (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_1, 0); -lean_dec(x_16); -x_17 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_1, 0, x_17); -x_18 = l_Lean_indentD(x_4); -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_1); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_1); -x_20 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_20, 0, x_2); -x_21 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_21, 0, x_20); -x_22 = l_Lean_indentD(x_4); -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); +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_7, 1); +lean_inc(x_16); +lean_dec(x_7); +x_17 = lean_st_ref_take(x_3, 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_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_21 = lean_ctor_get(x_18, 4); +x_22 = lean_apply_1(x_1, x_21); +lean_ctor_set(x_18, 4, x_22); +x_23 = lean_st_ref_set(x_3, x_18, x_19); +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_box(0); +lean_ctor_set(x_23, 0, x_26); return x_23; } -} else { -lean_object* x_24; -lean_dec(x_14); -x_24 = lean_box(0); -x_5 = x_24; -goto block_13; +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_23, 1); +lean_inc(x_27); +lean_dec(x_23); +x_28 = lean_box(0); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -lean_object* x_25; -x_25 = lean_box(0); -x_5 = x_25; -goto block_13; +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; 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_30 = lean_ctor_get(x_18, 0); +x_31 = lean_ctor_get(x_18, 1); +x_32 = lean_ctor_get(x_18, 2); +x_33 = lean_ctor_get(x_18, 3); +x_34 = lean_ctor_get(x_18, 4); +lean_inc(x_34); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_18); +x_35 = lean_apply_1(x_1, x_34); +x_36 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_36, 0, x_30); +lean_ctor_set(x_36, 1, x_31); +lean_ctor_set(x_36, 2, x_32); +lean_ctor_set(x_36, 3, x_33); +lean_ctor_set(x_36, 4, x_35); +x_37 = lean_st_ref_set(x_3, x_36, x_19); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_39 = x_37; +} else { + lean_dec_ref(x_37); + x_39 = lean_box(0); } -block_13: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -lean_dec(x_5); -x_6 = l_Lean_Meta_appendOptMessageData___closed__2; -x_7 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_7, 0, x_1); -lean_ctor_set(x_7, 1, x_6); -x_8 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_8, 0, x_2); -x_9 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_9, 0, x_8); -x_10 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_10, 0, x_7); -lean_ctor_set(x_10, 1, x_9); -x_11 = l_Lean_indentD(x_4); -x_12 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -return x_12; +x_40 = lean_box(0); +if (lean_is_scalar(x_39)) { + x_41 = lean_alloc_ctor(0, 2, 0); +} else { + x_41 = x_39; +} +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_38); +return x_41; } } } } -static lean_object* _init_l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Meta_modifyDiag___boxed(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 = l_Lean_warningAsError; -return x_1; +lean_object* x_7; +x_7 = l_Lean_Meta_modifyDiag(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_Lean_logAt___at_Lean_Meta_reportDiag___spec__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_EXPORT lean_object* l_Lean_Meta_recordUnfold(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_267; uint8_t x_268; -x_267 = 2; -x_268 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_267); -if (x_268 == 0) -{ -lean_object* x_269; -x_269 = lean_box(0); -x_9 = x_269; -goto block_266; -} -else +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -lean_object* x_270; uint8_t x_271; -lean_inc(x_2); -x_270 = l_Lean_MessageData_hasSyntheticSorry(x_2); -x_271 = lean_unbox(x_270); -lean_dec(x_270); -if (x_271 == 0) +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_272; -x_272 = lean_box(0); -x_9 = x_272; -goto block_266; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_273; lean_object* x_274; -lean_dec(x_2); -x_273 = lean_box(0); -x_274 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_274, 0, x_273); -lean_ctor_set(x_274, 1, x_8); -return x_274; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } -block_266: +else { -uint8_t x_10; lean_object* x_260; uint8_t x_261; uint8_t x_262; -lean_dec(x_9); -x_260 = lean_ctor_get(x_6, 2); -x_261 = 1; -x_262 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_261); -if (x_262 == 0) +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_16 = lean_ctor_get(x_7, 1); +lean_inc(x_16); +lean_dec(x_7); +x_17 = lean_st_ref_take(x_3, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 4); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) { -x_10 = x_3; -goto block_259; -} -else +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_18, 4); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_19, 0); +lean_inc(x_1); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_263; uint8_t x_264; -x_263 = l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2___closed__1; -x_264 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_260, x_263); -if (x_264 == 0) +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_unsigned_to_nat(1u); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +lean_ctor_set(x_19, 0, x_27); +x_28 = lean_st_ref_set(x_3, x_18, x_20); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) { -x_10 = x_3; -goto block_259; +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 { -uint8_t x_265; -x_265 = 2; -x_10 = x_265; -goto block_259; +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; } } -block_259: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; lean_object* x_19; -x_11 = lean_ctor_get(x_6, 0); -x_12 = lean_ctor_get(x_6, 1); -x_13 = lean_ctor_get(x_6, 5); -x_14 = lean_ctor_get(x_6, 6); -x_15 = lean_ctor_get(x_6, 7); -x_16 = l_Lean_replaceRef(x_1, x_13); -x_17 = 0; -x_18 = l_Lean_Syntax_getPos_x3f(x_16, x_17); -x_19 = l_Lean_Syntax_getTailPos_x3f(x_16, x_17); -if (lean_obj_tag(x_18) == 0) -{ -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; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; -x_20 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_FileMap_toPosition(x_12, x_23); -lean_inc(x_24); -x_25 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_24); -lean_inc(x_15); -lean_inc(x_14); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_14); -lean_ctor_set(x_26, 1, x_15); -x_27 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_21); -x_28 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_29 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_29, 0, x_11); -lean_ctor_set(x_29, 1, x_24); -lean_ctor_set(x_29, 2, x_25); -lean_ctor_set(x_29, 3, x_28); -lean_ctor_set(x_29, 4, x_27); -lean_ctor_set_uint8(x_29, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_29, sizeof(void*)*5 + 1, x_10); -x_30 = lean_st_ref_take(x_7, x_22); -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) +else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_34 = lean_ctor_get(x_31, 5); -x_35 = l_Lean_PersistentArray_push___rarg(x_34, x_29); -lean_ctor_set(x_31, 5, x_35); -x_36 = lean_st_ref_set(x_7, x_31, x_32); -x_37 = !lean_is_exclusive(x_36); -if (x_37 == 0) +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_35 = lean_ctor_get(x_25, 0); +lean_inc(x_35); +lean_dec(x_25); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_add(x_35, x_36); +lean_dec(x_35); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +lean_ctor_set(x_19, 0, x_38); +x_39 = lean_st_ref_set(x_3, x_18, x_20); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_38; lean_object* x_39; -x_38 = lean_ctor_get(x_36, 0); -lean_dec(x_38); -x_39 = lean_box(0); -lean_ctor_set(x_36, 0, x_39); -return x_36; +lean_object* x_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_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_36, 1); -lean_inc(x_40); -lean_dec(x_36); -x_41 = lean_box(0); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_object* x_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_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_31, 0); -x_44 = lean_ctor_get(x_31, 1); -x_45 = lean_ctor_get(x_31, 2); -x_46 = lean_ctor_get(x_31, 3); -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_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_19, 0); +x_47 = lean_ctor_get(x_19, 1); +x_48 = lean_ctor_get(x_19, 2); 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_31); -x_50 = l_Lean_PersistentArray_push___rarg(x_48, x_29); -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_46); -lean_ctor_set(x_51, 4, x_47); -lean_ctor_set(x_51, 5, x_50); -lean_ctor_set(x_51, 6, x_49); -x_52 = lean_st_ref_set(x_7, x_51, x_32); -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; +lean_dec(x_19); +lean_inc(x_1); +lean_inc(x_46); +x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_46, x_1); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_50 = lean_unsigned_to_nat(1u); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_50); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_47); +lean_ctor_set(x_52, 2, x_48); +lean_ctor_set(x_18, 4, x_52); +x_53 = lean_st_ref_set(x_3, x_18, x_20); +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_52); - x_54 = lean_box(0); + lean_dec_ref(x_53); + x_55 = lean_box(0); } -x_55 = lean_box(0); -if (lean_is_scalar(x_54)) { - x_56 = lean_alloc_ctor(0, 2, 0); +x_56 = lean_box(0); +if (lean_is_scalar(x_55)) { + x_57 = 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; + x_57 = x_55; } +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_54); +return x_57; } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_19); -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; 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_58 = lean_ctor_get(x_19, 0); -x_59 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_unsigned_to_nat(0u); -x_63 = l_Lean_FileMap_toPosition(x_12, x_62); -x_64 = l_Lean_FileMap_toPosition(x_12, x_58); +lean_object* x_58; lean_object* x_59; lean_object* x_60; 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_ctor_get(x_49, 0); +lean_inc(x_58); +lean_dec(x_49); +x_59 = lean_unsigned_to_nat(1u); +x_60 = lean_nat_add(x_58, x_59); lean_dec(x_58); -lean_ctor_set(x_19, 0, x_64); -lean_inc(x_15); -lean_inc(x_14); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_14); -lean_ctor_set(x_65, 1, x_15); -x_66 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_60); -x_67 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_68 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_68, 0, x_11); -lean_ctor_set(x_68, 1, x_63); -lean_ctor_set(x_68, 2, x_19); -lean_ctor_set(x_68, 3, x_67); -lean_ctor_set(x_68, 4, x_66); -lean_ctor_set_uint8(x_68, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_68, sizeof(void*)*5 + 1, x_10); -x_69 = lean_st_ref_take(x_7, x_61); -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_is_exclusive(x_70); -if (x_72 == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; uint8_t x_76; -x_73 = lean_ctor_get(x_70, 5); -x_74 = l_Lean_PersistentArray_push___rarg(x_73, x_68); -lean_ctor_set(x_70, 5, x_74); -x_75 = lean_st_ref_set(x_7, x_70, x_71); -x_76 = !lean_is_exclusive(x_75); -if (x_76 == 0) -{ -lean_object* x_77; lean_object* x_78; -x_77 = lean_ctor_get(x_75, 0); -lean_dec(x_77); -x_78 = lean_box(0); -lean_ctor_set(x_75, 0, x_78); -return x_75; +x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_46, x_1, x_60); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_61); +lean_ctor_set(x_62, 1, x_47); +lean_ctor_set(x_62, 2, x_48); +lean_ctor_set(x_18, 4, x_62); +x_63 = lean_st_ref_set(x_3, x_18, x_20); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_65 = x_63; +} else { + lean_dec_ref(x_63); + x_65 = lean_box(0); +} +x_66 = lean_box(0); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_65; +} +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; } -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_75, 1); -lean_inc(x_79); -lean_dec(x_75); -x_80 = lean_box(0); -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 { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_82 = lean_ctor_get(x_70, 0); -x_83 = lean_ctor_get(x_70, 1); -x_84 = lean_ctor_get(x_70, 2); -x_85 = lean_ctor_get(x_70, 3); -x_86 = lean_ctor_get(x_70, 4); -x_87 = lean_ctor_get(x_70, 5); -x_88 = lean_ctor_get(x_70, 6); -lean_inc(x_88); -lean_inc(x_87); -lean_inc(x_86); -lean_inc(x_85); -lean_inc(x_84); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_70); -x_89 = l_Lean_PersistentArray_push___rarg(x_87, x_68); -x_90 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_90, 0, x_82); -lean_ctor_set(x_90, 1, x_83); -lean_ctor_set(x_90, 2, x_84); -lean_ctor_set(x_90, 3, x_85); -lean_ctor_set(x_90, 4, x_86); -lean_ctor_set(x_90, 5, x_89); -lean_ctor_set(x_90, 6, x_88); -x_91 = lean_st_ref_set(x_7, x_90, x_71); -x_92 = lean_ctor_get(x_91, 1); -lean_inc(x_92); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_93 = x_91; +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_68 = lean_ctor_get(x_18, 0); +x_69 = lean_ctor_get(x_18, 1); +x_70 = lean_ctor_get(x_18, 2); +x_71 = lean_ctor_get(x_18, 3); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_18); +x_72 = lean_ctor_get(x_19, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_19, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_19, 2); +lean_inc(x_74); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + x_75 = x_19; } else { - lean_dec_ref(x_91); - x_93 = lean_box(0); + lean_dec_ref(x_19); + x_75 = lean_box(0); } -x_94 = lean_box(0); -if (lean_is_scalar(x_93)) { - x_95 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_1); +lean_inc(x_72); +x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_72, x_1); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_77 = lean_unsigned_to_nat(1u); +x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_72, x_1, x_77); +if (lean_is_scalar(x_75)) { + x_79 = lean_alloc_ctor(0, 3, 0); } else { - x_95 = x_93; + x_79 = x_75; } -lean_ctor_set(x_95, 0, x_94); -lean_ctor_set(x_95, 1, x_92); -return x_95; +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_73); +lean_ctor_set(x_79, 2, x_74); +x_80 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_80, 0, x_68); +lean_ctor_set(x_80, 1, x_69); +lean_ctor_set(x_80, 2, x_70); +lean_ctor_set(x_80, 3, x_71); +lean_ctor_set(x_80, 4, x_79); +x_81 = lean_st_ref_set(x_3, x_80, x_20); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); +} +x_84 = lean_box(0); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_83; } +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_96 = lean_ctor_get(x_19, 0); -lean_inc(x_96); -lean_dec(x_19); -x_97 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_98 = lean_ctor_get(x_97, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_97, 1); -lean_inc(x_99); -lean_dec(x_97); -x_100 = lean_unsigned_to_nat(0u); -x_101 = l_Lean_FileMap_toPosition(x_12, x_100); -x_102 = l_Lean_FileMap_toPosition(x_12, x_96); -lean_dec(x_96); -x_103 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_103, 0, x_102); -lean_inc(x_15); -lean_inc(x_14); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_14); -lean_ctor_set(x_104, 1, x_15); -x_105 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_98); -x_106 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_107 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_107, 0, x_11); -lean_ctor_set(x_107, 1, x_101); -lean_ctor_set(x_107, 2, x_103); -lean_ctor_set(x_107, 3, x_106); -lean_ctor_set(x_107, 4, x_105); -lean_ctor_set_uint8(x_107, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_107, sizeof(void*)*5 + 1, x_10); -x_108 = lean_st_ref_take(x_7, x_99); -x_109 = lean_ctor_get(x_108, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_108, 1); -lean_inc(x_110); -lean_dec(x_108); -x_111 = lean_ctor_get(x_109, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_109, 1); -lean_inc(x_112); -x_113 = lean_ctor_get(x_109, 2); -lean_inc(x_113); -x_114 = lean_ctor_get(x_109, 3); -lean_inc(x_114); -x_115 = lean_ctor_get(x_109, 4); -lean_inc(x_115); -x_116 = lean_ctor_get(x_109, 5); -lean_inc(x_116); -x_117 = lean_ctor_get(x_109, 6); -lean_inc(x_117); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - lean_ctor_release(x_109, 2); - lean_ctor_release(x_109, 3); - lean_ctor_release(x_109, 4); - lean_ctor_release(x_109, 5); - lean_ctor_release(x_109, 6); - x_118 = x_109; -} else { - lean_dec_ref(x_109); - x_118 = lean_box(0); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_86 = lean_ctor_get(x_76, 0); +lean_inc(x_86); +lean_dec(x_76); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_add(x_86, x_87); +lean_dec(x_86); +x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_72, x_1, x_88); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 3, 0); +} else { + x_90 = x_75; +} +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_73); +lean_ctor_set(x_90, 2, x_74); +x_91 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_91, 0, x_68); +lean_ctor_set(x_91, 1, x_69); +lean_ctor_set(x_91, 2, x_70); +lean_ctor_set(x_91, 3, x_71); +lean_ctor_set(x_91, 4, x_90); +x_92 = lean_st_ref_set(x_3, x_91, x_20); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; +} else { + lean_dec_ref(x_92); + x_94 = lean_box(0); } -x_119 = l_Lean_PersistentArray_push___rarg(x_116, x_107); -if (lean_is_scalar(x_118)) { - x_120 = lean_alloc_ctor(0, 7, 0); +x_95 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_118; + x_96 = x_94; } -lean_ctor_set(x_120, 0, x_111); -lean_ctor_set(x_120, 1, x_112); -lean_ctor_set(x_120, 2, x_113); -lean_ctor_set(x_120, 3, x_114); -lean_ctor_set(x_120, 4, x_115); -lean_ctor_set(x_120, 5, x_119); -lean_ctor_set(x_120, 6, x_117); -x_121 = lean_st_ref_set(x_7, x_120, x_110); -x_122 = lean_ctor_get(x_121, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_121)) { - lean_ctor_release(x_121, 0); - lean_ctor_release(x_121, 1); - x_123 = x_121; -} else { - lean_dec_ref(x_121); - x_123 = lean_box(0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } -x_124 = lean_box(0); -if (lean_is_scalar(x_123)) { - x_125 = lean_alloc_ctor(0, 2, 0); -} else { - x_125 = x_123; } -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_122); -return x_125; } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_recordUnfold___boxed(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_19) == 0) +lean_object* x_7; +x_7 = l_Lean_Meta_recordUnfold(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_Lean_Meta_recordDefEqHeuristic(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_126; -x_126 = !lean_is_exclusive(x_18); -if (x_126 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_127 = lean_ctor_get(x_18, 0); -x_128 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -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 = l_Lean_FileMap_toPosition(x_12, x_127); -lean_dec(x_127); -lean_inc(x_131); -lean_ctor_set(x_18, 0, x_131); -lean_inc(x_15); -lean_inc(x_14); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_14); -lean_ctor_set(x_132, 1, x_15); -x_133 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_129); -x_134 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_135 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_135, 0, x_11); -lean_ctor_set(x_135, 1, x_131); -lean_ctor_set(x_135, 2, x_18); -lean_ctor_set(x_135, 3, x_134); -lean_ctor_set(x_135, 4, x_133); -lean_ctor_set_uint8(x_135, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_135, sizeof(void*)*5 + 1, x_10); -x_136 = lean_st_ref_take(x_7, x_130); -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -lean_dec(x_136); -x_139 = !lean_is_exclusive(x_137); -if (x_139 == 0) -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; -x_140 = lean_ctor_get(x_137, 5); -x_141 = l_Lean_PersistentArray_push___rarg(x_140, x_135); -lean_ctor_set(x_137, 5, x_141); -x_142 = lean_st_ref_set(x_7, x_137, x_138); -x_143 = !lean_is_exclusive(x_142); -if (x_143 == 0) -{ -lean_object* x_144; lean_object* x_145; -x_144 = lean_ctor_get(x_142, 0); -lean_dec(x_144); -x_145 = lean_box(0); -lean_ctor_set(x_142, 0, x_145); -return x_142; +uint8_t x_10; +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_146 = lean_ctor_get(x_142, 1); -lean_inc(x_146); -lean_dec(x_142); -x_147 = lean_box(0); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_146); -return x_148; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_149 = lean_ctor_get(x_137, 0); -x_150 = lean_ctor_get(x_137, 1); -x_151 = lean_ctor_get(x_137, 2); -x_152 = lean_ctor_get(x_137, 3); -x_153 = lean_ctor_get(x_137, 4); -x_154 = lean_ctor_get(x_137, 5); -x_155 = lean_ctor_get(x_137, 6); -lean_inc(x_155); -lean_inc(x_154); -lean_inc(x_153); -lean_inc(x_152); -lean_inc(x_151); -lean_inc(x_150); -lean_inc(x_149); -lean_dec(x_137); -x_156 = l_Lean_PersistentArray_push___rarg(x_154, x_135); -x_157 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_157, 0, x_149); -lean_ctor_set(x_157, 1, x_150); -lean_ctor_set(x_157, 2, x_151); -lean_ctor_set(x_157, 3, x_152); -lean_ctor_set(x_157, 4, x_153); -lean_ctor_set(x_157, 5, x_156); -lean_ctor_set(x_157, 6, x_155); -x_158 = lean_st_ref_set(x_7, x_157, x_138); -x_159 = lean_ctor_get(x_158, 1); -lean_inc(x_159); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_160 = x_158; -} else { - lean_dec_ref(x_158); - x_160 = lean_box(0); -} -x_161 = lean_box(0); -if (lean_is_scalar(x_160)) { - x_162 = lean_alloc_ctor(0, 2, 0); -} else { - x_162 = x_160; -} -lean_ctor_set(x_162, 0, x_161); -lean_ctor_set(x_162, 1, x_159); -return x_162; -} +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_7, 1); +lean_inc(x_16); +lean_dec(x_7); +x_17 = lean_st_ref_take(x_3, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 4); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_18, 4); +lean_dec(x_22); +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_1); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_unsigned_to_nat(1u); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +lean_ctor_set(x_19, 1, x_27); +x_28 = lean_st_ref_set(x_3, x_18, x_20); +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_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_163 = lean_ctor_get(x_18, 0); -lean_inc(x_163); -lean_dec(x_18); -x_164 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_165 = lean_ctor_get(x_164, 0); -lean_inc(x_165); -x_166 = lean_ctor_get(x_164, 1); -lean_inc(x_166); -lean_dec(x_164); -x_167 = l_Lean_FileMap_toPosition(x_12, x_163); -lean_dec(x_163); -lean_inc(x_167); -x_168 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_168, 0, x_167); -lean_inc(x_15); -lean_inc(x_14); -x_169 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_169, 0, x_14); -lean_ctor_set(x_169, 1, x_15); -x_170 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_170, 0, x_169); -lean_ctor_set(x_170, 1, x_165); -x_171 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_172 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_172, 0, x_11); -lean_ctor_set(x_172, 1, x_167); -lean_ctor_set(x_172, 2, x_168); -lean_ctor_set(x_172, 3, x_171); -lean_ctor_set(x_172, 4, x_170); -lean_ctor_set_uint8(x_172, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_172, sizeof(void*)*5 + 1, x_10); -x_173 = lean_st_ref_take(x_7, x_166); -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_ctor_get(x_174, 0); -lean_inc(x_176); -x_177 = lean_ctor_get(x_174, 1); -lean_inc(x_177); -x_178 = lean_ctor_get(x_174, 2); -lean_inc(x_178); -x_179 = lean_ctor_get(x_174, 3); -lean_inc(x_179); -x_180 = lean_ctor_get(x_174, 4); -lean_inc(x_180); -x_181 = lean_ctor_get(x_174, 5); -lean_inc(x_181); -x_182 = lean_ctor_get(x_174, 6); -lean_inc(x_182); -if (lean_is_exclusive(x_174)) { - lean_ctor_release(x_174, 0); - lean_ctor_release(x_174, 1); - lean_ctor_release(x_174, 2); - lean_ctor_release(x_174, 3); - lean_ctor_release(x_174, 4); - lean_ctor_release(x_174, 5); - lean_ctor_release(x_174, 6); - x_183 = x_174; -} else { - lean_dec_ref(x_174); - x_183 = lean_box(0); -} -x_184 = l_Lean_PersistentArray_push___rarg(x_181, x_172); -if (lean_is_scalar(x_183)) { - x_185 = lean_alloc_ctor(0, 7, 0); -} else { - x_185 = x_183; -} -lean_ctor_set(x_185, 0, x_176); -lean_ctor_set(x_185, 1, x_177); -lean_ctor_set(x_185, 2, x_178); -lean_ctor_set(x_185, 3, x_179); -lean_ctor_set(x_185, 4, x_180); -lean_ctor_set(x_185, 5, x_184); -lean_ctor_set(x_185, 6, x_182); -x_186 = lean_st_ref_set(x_7, x_185, x_175); -x_187 = lean_ctor_get(x_186, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_186)) { - lean_ctor_release(x_186, 0); - lean_ctor_release(x_186, 1); - x_188 = x_186; -} else { - lean_dec_ref(x_186); - x_188 = lean_box(0); -} -x_189 = lean_box(0); -if (lean_is_scalar(x_188)) { - x_190 = lean_alloc_ctor(0, 2, 0); -} else { - x_190 = x_188; -} -lean_ctor_set(x_190, 0, x_189); -lean_ctor_set(x_190, 1, x_187); -return x_190; +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_191; uint8_t x_192; -x_191 = lean_ctor_get(x_18, 0); -lean_inc(x_191); -lean_dec(x_18); -x_192 = !lean_is_exclusive(x_19); -if (x_192 == 0) -{ -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; uint8_t x_206; -x_193 = lean_ctor_get(x_19, 0); -x_194 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_195 = lean_ctor_get(x_194, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_194, 1); -lean_inc(x_196); -lean_dec(x_194); -x_197 = l_Lean_FileMap_toPosition(x_12, x_191); -lean_dec(x_191); -x_198 = l_Lean_FileMap_toPosition(x_12, x_193); -lean_dec(x_193); -lean_ctor_set(x_19, 0, x_198); -lean_inc(x_15); -lean_inc(x_14); -x_199 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_199, 0, x_14); -lean_ctor_set(x_199, 1, x_15); -x_200 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_200, 0, x_199); -lean_ctor_set(x_200, 1, x_195); -x_201 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_202 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_202, 0, x_11); -lean_ctor_set(x_202, 1, x_197); -lean_ctor_set(x_202, 2, x_19); -lean_ctor_set(x_202, 3, x_201); -lean_ctor_set(x_202, 4, x_200); -lean_ctor_set_uint8(x_202, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_202, sizeof(void*)*5 + 1, x_10); -x_203 = lean_st_ref_take(x_7, x_196); -x_204 = lean_ctor_get(x_203, 0); -lean_inc(x_204); -x_205 = lean_ctor_get(x_203, 1); -lean_inc(x_205); -lean_dec(x_203); -x_206 = !lean_is_exclusive(x_204); -if (x_206 == 0) -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; uint8_t x_210; -x_207 = lean_ctor_get(x_204, 5); -x_208 = l_Lean_PersistentArray_push___rarg(x_207, x_202); -lean_ctor_set(x_204, 5, x_208); -x_209 = lean_st_ref_set(x_7, x_204, x_205); -x_210 = !lean_is_exclusive(x_209); -if (x_210 == 0) +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_35 = lean_ctor_get(x_25, 0); +lean_inc(x_35); +lean_dec(x_25); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_add(x_35, x_36); +lean_dec(x_35); +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +lean_ctor_set(x_19, 1, x_38); +x_39 = lean_st_ref_set(x_3, x_18, x_20); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_211; lean_object* x_212; -x_211 = lean_ctor_get(x_209, 0); -lean_dec(x_211); -x_212 = lean_box(0); -lean_ctor_set(x_209, 0, x_212); -return x_209; +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_213; lean_object* x_214; lean_object* x_215; -x_213 = lean_ctor_get(x_209, 1); -lean_inc(x_213); -lean_dec(x_209); -x_214 = lean_box(0); -x_215 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_215, 0, x_214); -lean_ctor_set(x_215, 1, x_213); -return x_215; +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_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_216 = lean_ctor_get(x_204, 0); -x_217 = lean_ctor_get(x_204, 1); -x_218 = lean_ctor_get(x_204, 2); -x_219 = lean_ctor_get(x_204, 3); -x_220 = lean_ctor_get(x_204, 4); -x_221 = lean_ctor_get(x_204, 5); -x_222 = lean_ctor_get(x_204, 6); -lean_inc(x_222); -lean_inc(x_221); -lean_inc(x_220); -lean_inc(x_219); -lean_inc(x_218); -lean_inc(x_217); -lean_inc(x_216); -lean_dec(x_204); -x_223 = l_Lean_PersistentArray_push___rarg(x_221, x_202); -x_224 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_224, 0, x_216); -lean_ctor_set(x_224, 1, x_217); -lean_ctor_set(x_224, 2, x_218); -lean_ctor_set(x_224, 3, x_219); -lean_ctor_set(x_224, 4, x_220); -lean_ctor_set(x_224, 5, x_223); -lean_ctor_set(x_224, 6, x_222); -x_225 = lean_st_ref_set(x_7, x_224, x_205); -x_226 = lean_ctor_get(x_225, 1); -lean_inc(x_226); -if (lean_is_exclusive(x_225)) { - lean_ctor_release(x_225, 0); - lean_ctor_release(x_225, 1); - x_227 = x_225; -} else { - lean_dec_ref(x_225); - x_227 = lean_box(0); -} -x_228 = lean_box(0); -if (lean_is_scalar(x_227)) { - x_229 = lean_alloc_ctor(0, 2, 0); -} else { - x_229 = x_227; -} -lean_ctor_set(x_229, 0, x_228); -lean_ctor_set(x_229, 1, x_226); -return x_229; -} -} -else -{ -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_230 = lean_ctor_get(x_19, 0); -lean_inc(x_230); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_19, 0); +x_47 = lean_ctor_get(x_19, 1); +x_48 = lean_ctor_get(x_19, 2); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); lean_dec(x_19); -x_231 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_2, x_4, x_5, x_6, x_7, x_8); -x_232 = lean_ctor_get(x_231, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_231, 1); -lean_inc(x_233); -lean_dec(x_231); -x_234 = l_Lean_FileMap_toPosition(x_12, x_191); -lean_dec(x_191); -x_235 = l_Lean_FileMap_toPosition(x_12, x_230); -lean_dec(x_230); -x_236 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_236, 0, x_235); -lean_inc(x_15); -lean_inc(x_14); -x_237 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_237, 0, x_14); -lean_ctor_set(x_237, 1, x_15); -x_238 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_238, 0, x_237); -lean_ctor_set(x_238, 1, x_232); -x_239 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; -lean_inc(x_11); -x_240 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_240, 0, x_11); -lean_ctor_set(x_240, 1, x_234); -lean_ctor_set(x_240, 2, x_236); -lean_ctor_set(x_240, 3, x_239); -lean_ctor_set(x_240, 4, x_238); -lean_ctor_set_uint8(x_240, sizeof(void*)*5, x_17); -lean_ctor_set_uint8(x_240, sizeof(void*)*5 + 1, x_10); -x_241 = lean_st_ref_take(x_7, x_233); -x_242 = lean_ctor_get(x_241, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -lean_dec(x_241); -x_244 = lean_ctor_get(x_242, 0); -lean_inc(x_244); -x_245 = lean_ctor_get(x_242, 1); -lean_inc(x_245); -x_246 = lean_ctor_get(x_242, 2); -lean_inc(x_246); -x_247 = lean_ctor_get(x_242, 3); -lean_inc(x_247); -x_248 = lean_ctor_get(x_242, 4); -lean_inc(x_248); -x_249 = lean_ctor_get(x_242, 5); -lean_inc(x_249); -x_250 = lean_ctor_get(x_242, 6); -lean_inc(x_250); -if (lean_is_exclusive(x_242)) { - lean_ctor_release(x_242, 0); - lean_ctor_release(x_242, 1); - lean_ctor_release(x_242, 2); - lean_ctor_release(x_242, 3); - lean_ctor_release(x_242, 4); - lean_ctor_release(x_242, 5); - lean_ctor_release(x_242, 6); - x_251 = x_242; -} else { - lean_dec_ref(x_242); - x_251 = lean_box(0); -} -x_252 = l_Lean_PersistentArray_push___rarg(x_249, x_240); -if (lean_is_scalar(x_251)) { - x_253 = lean_alloc_ctor(0, 7, 0); -} else { - x_253 = x_251; -} -lean_ctor_set(x_253, 0, x_244); -lean_ctor_set(x_253, 1, x_245); -lean_ctor_set(x_253, 2, x_246); -lean_ctor_set(x_253, 3, x_247); -lean_ctor_set(x_253, 4, x_248); -lean_ctor_set(x_253, 5, x_252); -lean_ctor_set(x_253, 6, x_250); -x_254 = lean_st_ref_set(x_7, x_253, x_243); -x_255 = lean_ctor_get(x_254, 1); -lean_inc(x_255); -if (lean_is_exclusive(x_254)) { - lean_ctor_release(x_254, 0); - lean_ctor_release(x_254, 1); - x_256 = x_254; -} else { - lean_dec_ref(x_254); - x_256 = lean_box(0); -} -x_257 = lean_box(0); -if (lean_is_scalar(x_256)) { - x_258 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_1); +lean_inc(x_47); +x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_47, x_1); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_50 = lean_unsigned_to_nat(1u); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_50); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_46); +lean_ctor_set(x_52, 1, x_51); +lean_ctor_set(x_52, 2, x_48); +lean_ctor_set(x_18, 4, x_52); +x_53 = lean_st_ref_set(x_3, x_18, x_20); +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 { - x_258 = x_256; + lean_dec_ref(x_53); + x_55 = lean_box(0); } -lean_ctor_set(x_258, 0, x_257); -lean_ctor_set(x_258, 1, x_255); -return x_258; +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; +x_58 = lean_ctor_get(x_49, 0); +lean_inc(x_58); +lean_dec(x_49); +x_59 = lean_unsigned_to_nat(1u); +x_60 = lean_nat_add(x_58, x_59); +lean_dec(x_58); +x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_47, x_1, x_60); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_46); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_48); +lean_ctor_set(x_18, 4, x_62); +x_63 = lean_st_ref_set(x_3, x_18, x_20); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_65 = x_63; +} else { + lean_dec_ref(x_63); + x_65 = lean_box(0); } +x_66 = lean_box(0); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_65; } +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; } } } -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_reportDiag___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) { -_start: +else { -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_5, 5); -lean_inc(x_8); -x_9 = l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2(x_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); -lean_dec(x_3); -lean_dec(x_8); -return x_9; -} +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_68 = lean_ctor_get(x_18, 0); +x_69 = lean_ctor_get(x_18, 1); +x_70 = lean_ctor_get(x_18, 2); +x_71 = lean_ctor_get(x_18, 3); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_18); +x_72 = lean_ctor_get(x_19, 0); +lean_inc(x_72); +x_73 = lean_ctor_get(x_19, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_19, 2); +lean_inc(x_74); +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + x_75 = x_19; +} else { + lean_dec_ref(x_19); + x_75 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__1() { -_start: +lean_inc(x_1); +lean_inc(x_73); +x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_73, x_1); +if (lean_obj_tag(x_76) == 0) { -lean_object* x_1; -x_1 = l_Lean_diag_threshold; -return x_1; +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_77 = lean_unsigned_to_nat(1u); +x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_77); +if (lean_is_scalar(x_75)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_75; } +lean_ctor_set(x_79, 0, x_72); +lean_ctor_set(x_79, 1, x_78); +lean_ctor_set(x_79, 2, x_74); +x_80 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_80, 0, x_68); +lean_ctor_set(x_80, 1, x_69); +lean_ctor_set(x_80, 2, x_70); +lean_ctor_set(x_80, 3, x_71); +lean_ctor_set(x_80, 4, x_79); +x_81 = lean_st_ref_set(x_3, x_80, x_20); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unfolded declarations:", 22); -return x_1; +x_84 = lean_box(0); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_83; } +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__3() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("used instances:", 15); -return x_1; +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_76, 0); +lean_inc(x_86); +lean_dec(x_76); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_add(x_86, x_87); +lean_dec(x_86); +x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_73, x_1, x_88); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 3, 0); +} else { + x_90 = x_75; } +lean_ctor_set(x_90, 0, x_72); +lean_ctor_set(x_90, 1, x_89); +lean_ctor_set(x_90, 2, x_74); +x_91 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_91, 0, x_68); +lean_ctor_set(x_91, 1, x_69); +lean_ctor_set(x_91, 2, x_70); +lean_ctor_set(x_91, 3, x_71); +lean_ctor_set(x_91, 4, x_90); +x_92 = lean_st_ref_set(x_3, x_91, x_20); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; +} else { + lean_dec_ref(x_92); + x_94 = lean_box(0); } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("`isDefEq` heuristic:", 20); -return x_1; +x_95 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(0, 2, 0); +} else { + x_96 = x_94; } +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("\nuse `set_option diag.threshould ` to control threshold for reporting counters", 83); -return x_1; } } -static lean_object* _init_l_Lean_Meta_reportDiag___closed__6() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_reportDiag___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_Lean_Meta_reportDiag___closed__7() { +LEAN_EXPORT lean_object* l_Lean_Meta_recordDefEqHeuristic___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_reportDiag___closed__6; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_7; +x_7 = l_Lean_Meta_recordDefEqHeuristic(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_Lean_Meta_reportDiag(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_Meta_recordInstance(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = l_Lean_isDiagnosticsEnabled(x_3, x_4, x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_unbox(x_7); -lean_dec(x_7); -if (x_8 == 0) +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) { -uint8_t x_9; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); +uint8_t x_10; lean_dec(x_1); -x_9 = !lean_is_exclusive(x_6); -if (x_9 == 0) +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_6, 0); -lean_dec(x_10); -x_11 = lean_box(0); -lean_ctor_set(x_6, 0, x_11); -return x_6; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_6, 1); -lean_inc(x_12); -lean_dec(x_6); -x_13 = lean_box(0); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; } } else { -lean_object* x_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_6, 1); -lean_inc(x_15); -lean_dec(x_6); -x_16 = lean_ctor_get(x_3, 2); +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_7, 1); lean_inc(x_16); -x_17 = l_Lean_Meta_reportDiag___closed__1; -x_18 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_16, x_17); -lean_dec(x_16); -x_19 = lean_st_ref_get(x_2, x_15); -x_20 = lean_ctor_get(x_19, 0); +lean_dec(x_7); +x_17 = lean_st_ref_take(x_3, x_16); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_18, 4); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); 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, 4); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); +lean_dec(x_17); +x_21 = !lean_is_exclusive(x_18); +if (x_21 == 0) +{ +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_18, 4); lean_dec(x_22); -lean_inc(x_18); -x_24 = l_Lean_Meta_mkMessageBodyFor_x3f(x_23, x_18); -x_25 = lean_st_ref_get(x_2, x_21); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); -lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_26, 4); -lean_inc(x_28); -lean_dec(x_26); -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); +x_23 = !lean_is_exclusive(x_19); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; +x_24 = lean_ctor_get(x_19, 2); +lean_inc(x_1); +lean_inc(x_24); +x_25 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_24, x_1); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_unsigned_to_nat(1u); +x_27 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_26); +lean_ctor_set(x_19, 2, x_27); +x_28 = lean_st_ref_set(x_3, x_18, x_20); +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); -lean_inc(x_18); -x_30 = l_Lean_Meta_mkMessageBodyFor_x3f(x_29, x_18); -x_31 = lean_st_ref_get(x_2, x_27); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +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_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_33 = lean_ctor_get(x_31, 0); -x_34 = lean_ctor_get(x_31, 1); -x_35 = lean_ctor_get(x_33, 4); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_35 = lean_ctor_get(x_25, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_ctor_get(x_35, 2); -lean_inc(x_36); +lean_dec(x_25); +x_36 = lean_unsigned_to_nat(1u); +x_37 = lean_nat_add(x_35, x_36); lean_dec(x_35); -x_37 = l_Lean_Meta_mkMessageBodyFor_x3f(x_36, x_18); -if (lean_obj_tag(x_24) == 0) -{ -if (lean_obj_tag(x_30) == 0) -{ -if (lean_obj_tag(x_37) == 0) +x_38 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_24, x_1, x_37); +lean_ctor_set(x_19, 2, x_38); +x_39 = lean_st_ref_set(x_3, x_18, x_20); +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) { -lean_object* x_38; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_38 = lean_box(0); -lean_ctor_set(x_31, 0, x_38); -return x_31; +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_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; -lean_free_object(x_31); -x_39 = l_Lean_MessageData_nil; -x_40 = l_Lean_Meta_reportDiag___closed__2; -x_41 = l_Lean_Meta_appendOptMessageData(x_39, x_40, x_24); -x_42 = l_Lean_Meta_reportDiag___closed__3; -x_43 = l_Lean_Meta_appendOptMessageData(x_41, x_42, x_37); -x_44 = l_Lean_Meta_reportDiag___closed__4; -x_45 = l_Lean_Meta_appendOptMessageData(x_43, x_44, x_30); -x_46 = l_Lean_Meta_reportDiag___closed__7; -x_47 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = 0; -x_49 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_47, x_48, x_1, x_2, x_3, x_4, x_34); -return x_49; +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_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; -lean_free_object(x_31); -x_50 = l_Lean_MessageData_nil; -x_51 = l_Lean_Meta_reportDiag___closed__2; -x_52 = l_Lean_Meta_appendOptMessageData(x_50, x_51, x_24); -x_53 = l_Lean_Meta_reportDiag___closed__3; -x_54 = l_Lean_Meta_appendOptMessageData(x_52, x_53, x_37); -x_55 = l_Lean_Meta_reportDiag___closed__4; -x_56 = l_Lean_Meta_appendOptMessageData(x_54, x_55, x_30); -x_57 = l_Lean_Meta_reportDiag___closed__7; -x_58 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_58, 1, x_57); -x_59 = 0; -x_60 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_58, x_59, x_1, x_2, x_3, x_4, x_34); -return x_60; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_ctor_get(x_19, 0); +x_47 = lean_ctor_get(x_19, 1); +x_48 = lean_ctor_get(x_19, 2); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_19); +lean_inc(x_1); +lean_inc(x_48); +x_49 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_48, x_1); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_50 = lean_unsigned_to_nat(1u); +x_51 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_50); +x_52 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_52, 0, x_46); +lean_ctor_set(x_52, 1, x_47); +lean_ctor_set(x_52, 2, x_51); +lean_ctor_set(x_18, 4, x_52); +x_53 = lean_st_ref_set(x_3, x_18, x_20); +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_61; 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_free_object(x_31); -x_61 = l_Lean_MessageData_nil; -x_62 = l_Lean_Meta_reportDiag___closed__2; -x_63 = l_Lean_Meta_appendOptMessageData(x_61, x_62, x_24); -x_64 = l_Lean_Meta_reportDiag___closed__3; -x_65 = l_Lean_Meta_appendOptMessageData(x_63, x_64, x_37); -x_66 = l_Lean_Meta_reportDiag___closed__4; -x_67 = l_Lean_Meta_appendOptMessageData(x_65, x_66, x_30); -x_68 = l_Lean_Meta_reportDiag___closed__7; -x_69 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_69, 0, x_67); -lean_ctor_set(x_69, 1, x_68); -x_70 = 0; -x_71 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_69, x_70, x_1, x_2, x_3, x_4, x_34); -return x_71; +lean_object* x_58; lean_object* x_59; lean_object* x_60; 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_ctor_get(x_49, 0); +lean_inc(x_58); +lean_dec(x_49); +x_59 = lean_unsigned_to_nat(1u); +x_60 = lean_nat_add(x_58, x_59); +lean_dec(x_58); +x_61 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_48, x_1, x_60); +x_62 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_62, 0, x_46); +lean_ctor_set(x_62, 1, x_47); +lean_ctor_set(x_62, 2, x_61); +lean_ctor_set(x_18, 4, x_62); +x_63 = lean_st_ref_set(x_3, x_18, x_20); +x_64 = lean_ctor_get(x_63, 1); +lean_inc(x_64); +if (lean_is_exclusive(x_63)) { + lean_ctor_release(x_63, 0); + lean_ctor_release(x_63, 1); + x_65 = x_63; +} else { + lean_dec_ref(x_63); + x_65 = lean_box(0); +} +x_66 = lean_box(0); +if (lean_is_scalar(x_65)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_65; +} +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_64); +return x_67; +} } } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_31, 0); -x_73 = lean_ctor_get(x_31, 1); -lean_inc(x_73); +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_68 = lean_ctor_get(x_18, 0); +x_69 = lean_ctor_get(x_18, 1); +x_70 = lean_ctor_get(x_18, 2); +x_71 = lean_ctor_get(x_18, 3); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_18); +x_72 = lean_ctor_get(x_19, 0); lean_inc(x_72); -lean_dec(x_31); -x_74 = lean_ctor_get(x_72, 4); +x_73 = lean_ctor_get(x_19, 1); +lean_inc(x_73); +x_74 = lean_ctor_get(x_19, 2); lean_inc(x_74); -lean_dec(x_72); -x_75 = lean_ctor_get(x_74, 2); -lean_inc(x_75); -lean_dec(x_74); -x_76 = l_Lean_Meta_mkMessageBodyFor_x3f(x_75, x_18); -if (lean_obj_tag(x_24) == 0) -{ -if (lean_obj_tag(x_30) == 0) -{ +if (lean_is_exclusive(x_19)) { + lean_ctor_release(x_19, 0); + lean_ctor_release(x_19, 1); + lean_ctor_release(x_19, 2); + x_75 = x_19; +} else { + lean_dec_ref(x_19); + x_75 = lean_box(0); +} +lean_inc(x_1); +lean_inc(x_74); +x_76 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Kernel_Diagnostics_recordUnfold___spec__1(x_74, x_1); if (lean_obj_tag(x_76) == 0) { -lean_object* x_77; lean_object* x_78; -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_77 = lean_box(0); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_73); -return x_78; +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_77 = lean_unsigned_to_nat(1u); +x_78 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_77); +if (lean_is_scalar(x_75)) { + x_79 = lean_alloc_ctor(0, 3, 0); +} else { + x_79 = x_75; } -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; lean_object* x_89; -x_79 = l_Lean_MessageData_nil; -x_80 = l_Lean_Meta_reportDiag___closed__2; -x_81 = l_Lean_Meta_appendOptMessageData(x_79, x_80, x_24); -x_82 = l_Lean_Meta_reportDiag___closed__3; -x_83 = l_Lean_Meta_appendOptMessageData(x_81, x_82, x_76); -x_84 = l_Lean_Meta_reportDiag___closed__4; -x_85 = l_Lean_Meta_appendOptMessageData(x_83, x_84, x_30); -x_86 = l_Lean_Meta_reportDiag___closed__7; -x_87 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -x_88 = 0; -x_89 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_87, x_88, x_1, x_2, x_3, x_4, x_73); -return x_89; +lean_ctor_set(x_79, 0, x_72); +lean_ctor_set(x_79, 1, x_73); +lean_ctor_set(x_79, 2, x_78); +x_80 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_80, 0, x_68); +lean_ctor_set(x_80, 1, x_69); +lean_ctor_set(x_80, 2, x_70); +lean_ctor_set(x_80, 3, x_71); +lean_ctor_set(x_80, 4, x_79); +x_81 = lean_st_ref_set(x_3, x_80, x_20); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +if (lean_is_exclusive(x_81)) { + lean_ctor_release(x_81, 0); + lean_ctor_release(x_81, 1); + x_83 = x_81; +} else { + lean_dec_ref(x_81); + x_83 = lean_box(0); +} +x_84 = lean_box(0); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); +} else { + x_85 = x_83; } +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } 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; uint8_t x_99; lean_object* x_100; -x_90 = l_Lean_MessageData_nil; -x_91 = l_Lean_Meta_reportDiag___closed__2; -x_92 = l_Lean_Meta_appendOptMessageData(x_90, x_91, x_24); -x_93 = l_Lean_Meta_reportDiag___closed__3; -x_94 = l_Lean_Meta_appendOptMessageData(x_92, x_93, x_76); -x_95 = l_Lean_Meta_reportDiag___closed__4; -x_96 = l_Lean_Meta_appendOptMessageData(x_94, x_95, x_30); -x_97 = l_Lean_Meta_reportDiag___closed__7; -x_98 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_98, 0, x_96); -lean_ctor_set(x_98, 1, x_97); -x_99 = 0; -x_100 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_98, x_99, x_1, x_2, x_3, x_4, x_73); -return x_100; +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_76, 0); +lean_inc(x_86); +lean_dec(x_76); +x_87 = lean_unsigned_to_nat(1u); +x_88 = lean_nat_add(x_86, x_87); +lean_dec(x_86); +x_89 = l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(x_74, x_1, x_88); +if (lean_is_scalar(x_75)) { + x_90 = lean_alloc_ctor(0, 3, 0); +} else { + x_90 = x_75; } +lean_ctor_set(x_90, 0, x_72); +lean_ctor_set(x_90, 1, x_73); +lean_ctor_set(x_90, 2, x_89); +x_91 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_91, 0, x_68); +lean_ctor_set(x_91, 1, x_69); +lean_ctor_set(x_91, 2, x_70); +lean_ctor_set(x_91, 3, x_71); +lean_ctor_set(x_91, 4, x_90); +x_92 = lean_st_ref_set(x_3, x_91, x_20); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_92)) { + lean_ctor_release(x_92, 0); + lean_ctor_release(x_92, 1); + x_94 = x_92; +} else { + lean_dec_ref(x_92); + x_94 = lean_box(0); } -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_object* x_111; -x_101 = l_Lean_MessageData_nil; -x_102 = l_Lean_Meta_reportDiag___closed__2; -x_103 = l_Lean_Meta_appendOptMessageData(x_101, x_102, x_24); -x_104 = l_Lean_Meta_reportDiag___closed__3; -x_105 = l_Lean_Meta_appendOptMessageData(x_103, x_104, x_76); -x_106 = l_Lean_Meta_reportDiag___closed__4; -x_107 = l_Lean_Meta_appendOptMessageData(x_105, x_106, x_30); -x_108 = l_Lean_Meta_reportDiag___closed__7; -x_109 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -x_110 = 0; -x_111 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_109, x_110, x_1, x_2, x_3, x_4, x_73); -return x_111; +x_95 = lean_box(0); +if (lean_is_scalar(x_94)) { + x_96 = lean_alloc_ctor(0, 2, 0); +} else { + x_96 = x_94; +} +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_93); +return x_96; } } } } } -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_reportDiag___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_Lean_Meta_recordInstance___boxed(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 = lean_unbox(x_3); -lean_dec(x_3); -x_10 = l_Lean_logAt___at_Lean_Meta_reportDiag___spec__2(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); +lean_object* x_7; +x_7 = l_Lean_Meta_recordInstance(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_1); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_reportDiag___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: -{ -uint8_t x_8; lean_object* x_9; -x_8 = lean_unbox(x_2); +lean_dec(x_3); lean_dec(x_2); -x_9 = l_Lean_log___at_Lean_Meta_reportDiag___spec__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); -return x_9; +return x_7; } } LEAN_EXPORT lean_object* l_Lean_Meta_getLocalInstances(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -36146,21 +34670,30 @@ return x_4; static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1() { _start: { +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_instMetaEvalMetaM___rarg___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes(" =\?= ", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; +x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4() { _start: { lean_object* x_1; @@ -36168,16 +34701,16 @@ x_1 = lean_mk_string_from_bytes("\nwhile trying to unify", 22); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3; +x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6() { _start: { lean_object* x_1; @@ -36185,16 +34718,16 @@ x_1 = lean_mk_string_from_bytes("\nwith", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; +x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__8() { _start: { lean_object* x_1; @@ -36202,11 +34735,11 @@ x_1 = lean_mk_string_from_bytes(" : ", 3); return x_1; } } -static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__8() { +static lean_object* _init_l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7; +x_1 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__8; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -36270,7 +34803,7 @@ x_83 = lean_ctor_get(x_81, 1); lean_inc(x_83); lean_dec(x_81); x_84 = l_Lean_stringToMessageData(x_4); -x_85 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; +x_85 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; x_86 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_86, 0, x_85); lean_ctor_set(x_86, 1, x_84); @@ -36281,7 +34814,7 @@ x_88 = l_Lean_MessageData_ofLevel(x_5); x_89 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_89, 0, x_85); lean_ctor_set(x_89, 1, x_88); -x_90 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2; +x_90 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3; x_91 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_91, 0, x_89); lean_ctor_set(x_91, 1, x_90); @@ -36296,7 +34829,7 @@ x_95 = l_Lean_indentD(x_94); x_96 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_96, 0, x_87); lean_ctor_set(x_96, 1, x_95); -x_97 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4; +x_97 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; x_98 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_98, 0, x_96); lean_ctor_set(x_98, 1, x_97); @@ -36304,7 +34837,7 @@ x_99 = l_Lean_MessageData_ofExpr(x_15); x_100 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_100, 0, x_85); lean_ctor_set(x_100, 1, x_99); -x_101 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__8; +x_101 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__9; x_102 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_102, 0, x_100); lean_ctor_set(x_102, 1, x_101); @@ -36319,7 +34852,7 @@ x_106 = l_Lean_indentD(x_105); x_107 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_107, 0, x_98); lean_ctor_set(x_107, 1, x_106); -x_108 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6; +x_108 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7; x_109 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_109, 0, x_107); lean_ctor_set(x_109, 1, x_108); @@ -36387,7 +34920,7 @@ lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean lean_dec(x_22); lean_dec(x_20); x_25 = l_Lean_stringToMessageData(x_4); -x_26 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; +x_26 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; x_27 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_27, 0, x_26); lean_ctor_set(x_27, 1, x_25); @@ -36398,7 +34931,7 @@ x_29 = l_Lean_MessageData_ofLevel(x_5); x_30 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_30, 0, x_26); lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2; +x_31 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3; x_32 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_32, 0, x_30); lean_ctor_set(x_32, 1, x_31); @@ -36413,7 +34946,7 @@ x_36 = l_Lean_indentD(x_35); x_37 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_37, 0, x_28); lean_ctor_set(x_37, 1, x_36); -x_38 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4; +x_38 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; x_39 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_39, 0, x_37); lean_ctor_set(x_39, 1, x_38); @@ -36422,7 +34955,7 @@ x_41 = l_Lean_indentD(x_40); x_42 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_42, 0, x_39); lean_ctor_set(x_42, 1, x_41); -x_43 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6; +x_43 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7; x_44 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); @@ -36472,7 +35005,7 @@ lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean lean_dec(x_22); lean_dec(x_20); x_52 = l_Lean_stringToMessageData(x_4); -x_53 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; +x_53 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; x_54 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_54, 0, x_53); lean_ctor_set(x_54, 1, x_52); @@ -36483,7 +35016,7 @@ x_56 = l_Lean_MessageData_ofLevel(x_5); x_57 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_57, 0, x_53); lean_ctor_set(x_57, 1, x_56); -x_58 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2; +x_58 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3; x_59 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_59, 0, x_57); lean_ctor_set(x_59, 1, x_58); @@ -36498,7 +35031,7 @@ x_63 = l_Lean_indentD(x_62); x_64 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_64, 0, x_55); lean_ctor_set(x_64, 1, x_63); -x_65 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__4; +x_65 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__5; x_66 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_66, 0, x_64); lean_ctor_set(x_66, 1, x_65); @@ -36507,7 +35040,7 @@ x_68 = l_Lean_indentD(x_67); x_69 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_69, 0, x_66); lean_ctor_set(x_69, 1, x_68); -x_70 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__6; +x_70 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__7; x_71 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_71, 0, x_69); lean_ctor_set(x_71, 1, x_70); @@ -36545,7 +35078,7 @@ lean_dec(x_4); lean_dec(x_3); x_9 = l_Lean_stringToMessageData(x_1); lean_dec(x_1); -x_10 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkMessageBodyFor_x3f___spec__1___closed__5; +x_10 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__1; x_11 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -36558,7 +35091,7 @@ x_14 = l_Lean_MessageData_ofLevel(x_13); x_15 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_15, 0, x_10); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__2; +x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___lambda__3___closed__3; x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -39705,119 +38238,7 @@ x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_4; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; double x_18; double x_19; double x_20; double x_21; double x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_15 = lean_ctor_get(x_13, 0); -x_16 = 0; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Float_ofScientific(x_8, x_16, x_17); -lean_dec(x_8); -x_19 = l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; -x_20 = lean_float_div(x_18, x_19); -x_21 = l_Float_ofScientific(x_15, x_16, x_17); -lean_dec(x_15); -x_22 = lean_float_div(x_21, x_19); -x_23 = lean_box_float(x_20); -x_24 = lean_box_float(x_22); -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, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_13, 0, x_26); -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_dec(x_13); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_8, x_29, x_30); -lean_dec(x_8); -x_32 = l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = l_Float_ofScientific(x_27, x_29, x_30); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_32); -x_36 = lean_box_float(x_33); -x_37 = lean_box_float(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_11); -lean_ctor_set(x_39, 1, 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_28); -return x_40; -} -} -else -{ -uint8_t x_41; -lean_dec(x_8); -x_41 = !lean_is_exclusive(x_10); -if (x_41 == 0) -{ -return x_10; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_10, 0); -x_43 = lean_ctor_get(x_10, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_10); -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_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, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; @@ -40086,7 +38507,7 @@ return x_99; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { if (lean_obj_tag(x_1) == 0) @@ -40111,89 +38532,23 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_6); lean_inc(x_9); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_4, x_7, x_8, x_9, x_10, x_13); lean_dec(x_9); lean_dec(x_4); return x_14; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -40201,7 +38556,7 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_processPostponed___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_Meta_processPostponed___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, uint8_t x_8, double x_9, double 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: { double x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; @@ -40214,7 +38569,7 @@ lean_ctor_set(x_18, 1, x_3); lean_ctor_set_float(x_18, sizeof(void*)*2, x_17); lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_17); lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_2); -x_19 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1; +x_19 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2___closed__1; x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_19); if (x_20 == 0) { @@ -40224,7 +38579,7 @@ lean_object* x_21; lean_object* x_22; lean_dec(x_3); lean_dec(x_1); x_21 = lean_box(0); -x_22 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); return x_22; } else @@ -40238,7 +38593,7 @@ lean_ctor_set_float(x_23, sizeof(void*)*2, x_9); lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_2); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); +x_25 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); return x_25; } } @@ -40253,12 +38608,12 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); return x_28; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_processPostponed___spec__1___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -40266,16 +38621,16 @@ x_1 = lean_mk_string_from_bytes(" +#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 +extern lean_object* l_Lean_diagExt; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_subCounters(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_appendSection___closed__1; +static lean_object* l_Lean_Meta_reportDiag___closed__11; +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold(lean_object*); +lean_object* lean_array_push(lean_object*, lean_object*); +lean_object* l_Lean_Name_lt___boxed(lean_object*, lean_object*); +lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkDiagSummary___closed__3; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiagSummary_isEmpty___boxed(lean_object*); +lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4___rarg(lean_object*, lean_object*, lean_object*); +lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_instInhabitedDiagSummary___closed__1; +LEAN_EXPORT uint8_t l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1(lean_object*); +LEAN_EXPORT uint8_t l_Lean_Meta_DiagSummary_isEmpty(lean_object*); +lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Meta_registerCoercion___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_subCounters___rarg___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_subCounters___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__17; +static lean_object* l_Lean_Meta_reportDiag___closed__3; +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiagSummary_max___default; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1; +LEAN_EXPORT uint8_t l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_appendSection___closed__3; +LEAN_EXPORT uint8_t l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1(lean_object*, uint8_t, lean_object*); +extern lean_object* l_instInhabitedNat; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_instInhabitedDiagSummary; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_isInstanceCore(lean_object*, lean_object*); +lean_object* l_Lean_EnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +uint8_t lean_get_reducibility_status(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummary(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_DiagSummary_data___default; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__14; +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1___boxed(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__2; +static lean_object* l_Lean_Meta_mkDiagSummary___closed__2; +static lean_object* l_Lean_Meta_reportDiag___closed__6; +static double l_Lean_Meta_appendSection___closed__4; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___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_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(lean_object*); +extern lean_object* l_Lean_MessageData_nil; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5; +lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___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*); +static lean_object* l_Lean_Meta_reportDiag___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_appendSection(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__5; +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__8; +static lean_object* l_Lean_Meta_subCounters___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_appendSection___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_Kernel_instInhabitedDiagnostics; +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummary___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__4; +static lean_object* l_Lean_Meta_reportDiag___closed__15; +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__12; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkDiagSummary___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_mkDiagSummary___spec__3(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__7; +static lean_object* l_Lean_Meta_reportDiag___closed__8; +extern lean_object* l_Lean_diagnostics_threshold; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9; +static lean_object* l_Lean_Meta_reportDiag___closed__13; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3; +size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_instInhabitedName; +lean_object* lean_array_uget(lean_object*, size_t); +static lean_object* l_Lean_Meta_appendSection___closed__2; +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1(lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +lean_object* lean_array_get_size(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_subCounters___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___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_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__10; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2(lean_object*); +uint8_t lean_usize_dec_lt(size_t, size_t); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(lean_object*, lean_object*, lean_object*); +lean_object* lean_nat_add(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__16; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1; +lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_reportDiag___closed__1; +static lean_object* l_Lean_Meta_reportDiag___closed__12; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +lean_object* l_Lean_PersistentHashMap_insert___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___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; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_array_uget(x_2, x_3); +switch (lean_obj_tag(x_7)) { +case 0: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_1); +x_10 = lean_apply_3(x_1, x_5, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; size_t x_15; size_t x_16; +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_14; +goto _start; +} +} +case 1: +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +lean_dec(x_7); +lean_inc(x_1); +x_19 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(x_1, x_18, x_5); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; size_t x_24; size_t x_25; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); +x_24 = 1; +x_25 = lean_usize_add(x_3, x_24); +x_3 = x_25; +x_5 = x_23; +goto _start; +} +} +default: +{ +size_t x_27; size_t x_28; +x_27 = 1; +x_28 = lean_usize_add(x_3, x_27); +x_3 = x_28; +goto _start; +} +} +} +else +{ +lean_object* x_30; +lean_dec(x_1); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_5); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg___boxed), 5, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_5); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_6); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_fget(x_2, x_5); +x_11 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +x_12 = lean_apply_3(x_1, x_6, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_5); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_18; +x_6 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed), 6, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_3); +return x_8; +} +else +{ +uint8_t x_9; +x_9 = lean_nat_dec_le(x_5, x_5); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(x_1, x_4, x_11, x_12, x_3); +lean_dec(x_4); +return x_13; +} +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(x_1, x_14, x_15, lean_box(0), x_16, x_3); +lean_dec(x_15); +lean_dec(x_14); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg), 3, 0); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_collectAboveThreshold___spec__3___rarg(x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_apply_2(x_1, x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___lambda__1), 4, 1); +lean_closure_set(x_6, 0, x_5); +x_7 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___rarg(x_3, x_6, x_4); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___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; lean_object* x_7; uint8_t x_8; +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_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); +lean_dec(x_3); +x_8 = lean_nat_dec_eq(x_5, x_7); +if (x_8 == 0) +{ +uint8_t x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_9 = lean_nat_dec_lt(x_7, x_5); +lean_dec(x_5); +lean_dec(x_7); +x_10 = lean_box(x_9); +return x_10; +} +else +{ +lean_object* x_11; +lean_dec(x_7); +lean_dec(x_5); +x_11 = lean_apply_2(x_1, x_4, x_6); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___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; +lean_inc(x_1); +x_5 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg___lambda__1), 3, 1); +lean_closure_set(x_5, 0, x_1); +x_6 = lean_nat_dec_lt(x_3, x_4); +if (x_6 == 0) +{ +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_2; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_inc(x_3); +x_7 = l_Array_qpartition___rarg(x_2, x_5, x_3, x_4); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_nat_dec_le(x_4, x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_1); +x_11 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg(x_1, x_9, x_3, x_8); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_8, x_12); +lean_dec(x_8); +x_2 = x_11; +x_3 = x_13; +goto _start; +} +else +{ +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg___boxed), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_nat_dec_lt(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +lean_inc(x_6); +x_10 = lean_apply_1(x_2, x_6); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_4); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_array_push(x_4, x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_3, 0); +x_16 = lean_ctor_get(x_3, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_3); +x_17 = lean_nat_dec_lt(x_1, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_2); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_4); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; +lean_inc(x_15); +x_19 = lean_apply_1(x_2, x_15); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_15); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_4); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_15); +lean_ctor_set(x_22, 1, x_16); +x_23 = lean_array_push(x_4, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_collectAboveThreshold___rarg___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_7 = lean_alloc_closure((void*)(l_Lean_Meta_collectAboveThreshold___rarg___lambda__1___boxed), 4, 2); +lean_closure_set(x_7, 0, x_4); +lean_closure_set(x_7, 1, x_5); +x_8 = l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +x_9 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg(x_1, x_2, x_3, x_8, x_7); +lean_dec(x_2); +lean_dec(x_1); +x_10 = lean_array_get_size(x_9); +x_11 = lean_unsigned_to_nat(1u); +x_12 = lean_nat_sub(x_10, x_11); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg(x_6, x_9, x_13, x_12); +lean_dec(x_12); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_collectAboveThreshold___rarg), 6, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___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; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_collectAboveThreshold___spec__4___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_collectAboveThreshold___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_collectAboveThreshold___spec__2(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___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_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Array_qsort_sort___at_Lean_Meta_collectAboveThreshold___spec__6___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___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_Lean_Meta_collectAboveThreshold___rarg___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, size_t x_8, size_t x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_eq(x_8, x_9); +if (x_11 == 0) +{ +lean_object* x_12; +x_12 = lean_array_uget(x_7, x_8); +switch (lean_obj_tag(x_12)) { +case 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); +lean_inc(x_6); +x_15 = lean_apply_3(x_6, x_10, x_13, x_14); +if (lean_obj_tag(x_15) == 0) +{ +uint8_t x_16; +lean_dec(x_6); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +return x_15; +} +else +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +lean_inc(x_17); +lean_dec(x_15); +x_18 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_18, 0, x_17); +return x_18; +} +} +else +{ +lean_object* x_19; size_t x_20; size_t x_21; +x_19 = lean_ctor_get(x_15, 0); +lean_inc(x_19); +lean_dec(x_15); +x_20 = 1; +x_21 = lean_usize_add(x_8, x_20); +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_box(0); +x_8 = x_21; +x_10 = x_19; +goto _start; +} +} +case 1: +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_12, 0); +lean_inc(x_23); +lean_dec(x_12); +lean_inc(x_6); +x_24 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg(x_1, x_2, lean_box(0), lean_box(0), lean_box(0), x_6, x_23, x_10); +if (lean_obj_tag(x_24) == 0) +{ +uint8_t x_25; +lean_dec(x_6); +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +return x_24; +} +else +{ +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_27, 0, x_26); +return x_27; +} +} +else +{ +lean_object* x_28; size_t x_29; size_t x_30; +x_28 = lean_ctor_get(x_24, 0); +lean_inc(x_28); +lean_dec(x_24); +x_29 = 1; +x_30 = lean_usize_add(x_8, x_29); +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_box(0); +x_8 = x_30; +x_10 = x_28; +goto _start; +} +} +default: +{ +size_t x_32; size_t x_33; +x_32 = 1; +x_33 = lean_usize_add(x_8, x_32); +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_box(0); +x_8 = x_33; +goto _start; +} +} +} +else +{ +lean_object* x_35; +lean_dec(x_6); +x_35 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_35, 0, x_10); +return x_35; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg___boxed), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_array_get_size(x_7); +x_13 = lean_nat_dec_lt(x_10, x_12); +lean_dec(x_12); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_10); +lean_dec(x_6); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_11); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_array_fget(x_7, x_10); +x_16 = lean_array_fget(x_8, x_10); +lean_inc(x_6); +x_17 = lean_apply_3(x_6, x_11, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_18; +lean_dec(x_10); +lean_dec(x_6); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +return x_17; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_20, 0, x_19); +return x_20; +} +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +lean_inc(x_21); +lean_dec(x_17); +x_22 = lean_unsigned_to_nat(1u); +x_23 = lean_nat_add(x_10, x_22); +lean_dec(x_10); +x_3 = lean_box(0); +x_4 = lean_box(0); +x_5 = lean_box(0); +x_9 = lean_box(0); +x_10 = x_23; +x_11 = x_21; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5___rarg___boxed), 11, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___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) { +_start: +{ +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_array_get_size(x_9); +x_11 = lean_unsigned_to_nat(0u); +x_12 = lean_nat_dec_lt(x_11, x_10); +if (x_12 == 0) +{ +lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_8); +return x_13; +} +else +{ +uint8_t x_14; +x_14 = lean_nat_dec_le(x_10, x_10); +if (x_14 == 0) +{ +lean_object* x_15; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_8); +return x_15; +} +else +{ +size_t x_16; size_t x_17; lean_object* x_18; +x_16 = 0; +x_17 = lean_usize_of_nat(x_10); +lean_dec(x_10); +x_18 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg(x_1, x_2, lean_box(0), lean_box(0), lean_box(0), x_6, x_9, x_16, x_17, x_8); +lean_dec(x_9); +return x_18; +} +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_7, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_7, 1); +lean_inc(x_20); +lean_dec(x_7); +x_21 = lean_unsigned_to_nat(0u); +x_22 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5___rarg(x_1, x_2, lean_box(0), lean_box(0), lean_box(0), x_6, x_19, x_20, lean_box(0), x_21, x_8); +lean_dec(x_20); +lean_dec(x_19); +return x_22; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg___boxed), 8, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___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; +x_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +lean_dec(x_3); +x_7 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg(x_1, x_2, lean_box(0), lean_box(0), lean_box(0), x_4, x_6, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_collectAboveThreshold___spec__1___rarg___lambda__1), 4, 1); +lean_closure_set(x_6, 0, x_5); +x_7 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg(x_1, x_2, x_3, x_6, x_4); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +lean_dec(x_7); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg___boxed), 5, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_subCounters___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_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); +lean_dec(x_4); +lean_inc(x_6); +lean_inc(x_2); +lean_inc(x_1); +x_8 = l_Lean_PersistentHashMap_find_x3f___rarg(x_1, x_2, x_3, x_6); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = l_Lean_PersistentHashMap_insert___rarg(x_1, x_2, x_5, x_6, x_7); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_8, 0); +lean_inc(x_11); +lean_dec(x_8); +x_12 = lean_nat_sub(x_7, x_11); +lean_dec(x_11); +lean_dec(x_7); +x_13 = l_Lean_PersistentHashMap_insert___rarg(x_1, x_2, x_5, x_6, x_12); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +static lean_object* _init_l_Lean_Meta_subCounters___rarg___closed__1() { +_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_subCounters___rarg___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_subCounters___rarg___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_Meta_subCounters___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +x_5 = l_Lean_Meta_subCounters___rarg___closed__2; +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_7, 0, x_5); +lean_ctor_set(x_7, 1, x_6); +lean_inc(x_2); +lean_inc(x_1); +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_subCounters___rarg___lambda__1), 5, 3); +lean_closure_set(x_8, 0, x_1); +lean_closure_set(x_8, 1, x_2); +lean_closure_set(x_8, 2, x_4); +x_9 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg(x_1, x_2, x_3, x_7, x_8); +lean_dec(x_2); +lean_dec(x_1); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_subCounters(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_subCounters___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = lean_unbox_usize(x_8); +lean_dec(x_8); +x_12 = lean_unbox_usize(x_9); +lean_dec(x_9); +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_subCounters___spec__4___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11, x_12, x_10); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; +x_12 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_subCounters___spec__5___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_2); +lean_dec(x_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_subCounters___spec__3___rarg(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_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg___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_PersistentHashMap_foldlM___at_Lean_Meta_subCounters___spec__2___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg___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_PersistentHashMap_forIn___at_Lean_Meta_subCounters___spec__1___rarg(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_DiagSummary_data___default() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_DiagSummary_max___default() { +_start: +{ +lean_object* x_1; +x_1 = lean_unsigned_to_nat(0u); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_instInhabitedDiagSummary___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +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_instInhabitedDiagSummary() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Meta_instInhabitedDiagSummary___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_DiagSummary_isEmpty(lean_object* x_1) { +_start: +{ +lean_object* x_2; uint8_t x_3; +x_2 = lean_ctor_get(x_1, 0); +x_3 = l_Array_isEmpty___rarg(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_DiagSummary_isEmpty___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_DiagSummary_isEmpty(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___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; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_array_uget(x_2, x_3); +switch (lean_obj_tag(x_7)) { +case 0: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_1); +x_10 = lean_apply_3(x_1, x_5, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; size_t x_15; size_t x_16; +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_14; +goto _start; +} +} +case 1: +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +lean_dec(x_7); +lean_inc(x_1); +x_19 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4___rarg(x_1, x_18, x_5); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; size_t x_24; size_t x_25; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); +x_24 = 1; +x_25 = lean_usize_add(x_3, x_24); +x_3 = x_25; +x_5 = x_23; +goto _start; +} +} +default: +{ +size_t x_27; size_t x_28; +x_27 = 1; +x_28 = lean_usize_add(x_3, x_27); +x_3 = x_28; +goto _start; +} +} +} +else +{ +lean_object* x_30; +lean_dec(x_1); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_5); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___rarg___boxed), 5, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_5); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_6); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_fget(x_2, x_5); +x_11 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +x_12 = lean_apply_3(x_1, x_6, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_5); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_18; +x_6 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg___boxed), 6, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_3); +return x_8; +} +else +{ +uint8_t x_9; +x_9 = lean_nat_dec_le(x_5, x_5); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___rarg(x_1, x_4, x_11, x_12, x_3); +lean_dec(x_4); +return x_13; +} +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg(x_1, x_14, x_15, lean_box(0), x_16, x_3); +lean_dec(x_15); +lean_dec(x_14); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_mkDiagSummary___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_mkDiagSummary___spec__4___rarg(x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_apply_2(x_1, x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2(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_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2___lambda__1), 4, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_mkDiagSummary___spec__3(x_1, x_4, x_2); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7___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; lean_object* x_7; uint8_t x_8; +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_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); +lean_dec(x_3); +x_8 = lean_nat_dec_eq(x_5, x_7); +if (x_8 == 0) +{ +uint8_t x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_9 = lean_nat_dec_lt(x_7, x_5); +lean_dec(x_5); +lean_dec(x_7); +x_10 = lean_box(x_9); +return x_10; +} +else +{ +lean_object* x_11; +lean_dec(x_7); +lean_dec(x_5); +x_11 = lean_apply_2(x_1, x_4, x_6); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7(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; +lean_inc(x_1); +x_5 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7___lambda__1), 3, 1); +lean_closure_set(x_5, 0, x_1); +x_6 = lean_nat_dec_lt(x_3, x_4); +if (x_6 == 0) +{ +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_2; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_inc(x_3); +x_7 = l_Array_qpartition___rarg(x_2, x_5, x_3, x_4); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_nat_dec_le(x_4, x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_1); +x_11 = l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7(x_1, x_9, x_3, x_8); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_8, x_12); +lean_dec(x_8); +x_2 = x_11; +x_3 = x_13; +goto _start; +} +else +{ +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_nat_dec_lt(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +lean_inc(x_6); +x_10 = lean_apply_1(x_2, x_6); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_4); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_array_push(x_4, x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_3, 0); +x_16 = lean_ctor_get(x_3, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_3); +x_17 = lean_nat_dec_lt(x_1, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_2); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_4); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; +lean_inc(x_15); +x_19 = lean_apply_1(x_2, x_15); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_15); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_4); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_15); +lean_ctor_set(x_22, 1, x_16); +x_23 = lean_array_push(x_4, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_alloc_closure((void*)(l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1___boxed), 4, 2); +lean_closure_set(x_5, 0, x_2); +lean_closure_set(x_5, 1, x_3); +x_6 = l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +x_7 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_mkDiagSummary___spec__2(x_1, x_6, x_5); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_8, x_9); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7(x_4, x_7, x_11, x_10); +lean_dec(x_10); +return x_12; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ↦ ", 5); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\n", 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6; +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_Meta_mkDiagSummary___spec__8___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ", 2); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10; +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_Meta_mkDiagSummary___spec__8___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; +x_10 = lean_usize_dec_lt(x_3, x_2); +if (x_10 == 0) +{ +lean_object* x_11; +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_1, x_3); +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_mkConstWithLevelParams___at_Lean_Meta_registerCoercion___spec__1(x_13, 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; uint8_t x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +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_isEmpty___rarg(x_4); +x_19 = l_Lean_MessageData_ofConst(x_16); +x_20 = l___private_Init_Data_Repr_0__Nat_reprFast(x_14); +x_21 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_21, 0, x_20); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +if (x_18 == 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; size_t x_31; size_t x_32; +x_23 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__8; +x_24 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_19); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_22); +x_28 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___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_array_push(x_4, x_29); +x_31 = 1; +x_32 = lean_usize_add(x_3, x_31); +x_3 = x_32; +x_4 = x_30; +x_9 = x_17; +goto _start; +} +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; size_t x_42; size_t x_43; +x_34 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__12; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_19); +x_36 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4; +x_37 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_22); +x_39 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_array_push(x_4, x_40); +x_42 = 1; +x_43 = lean_usize_add(x_3, x_42); +x_3 = x_43; +x_4 = x_41; +x_9 = x_17; +goto _start; +} +} +else +{ +uint8_t x_45; +lean_dec(x_14); +lean_dec(x_4); +x_45 = !lean_is_exclusive(x_15); +if (x_45 == 0) +{ +return x_15; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_15, 0); +x_47 = lean_ctor_get(x_15, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_15); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_mkDiagSummary___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics_threshold; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_mkDiagSummary___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Name_lt___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_mkDiagSummary___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_instInhabitedName; +x_2 = l_instInhabitedNat; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummary(lean_object* x_1, 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; uint8_t x_13; +x_8 = lean_ctor_get(x_5, 2); +x_9 = l_Lean_Meta_mkDiagSummary___closed__1; +x_10 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_8, x_9); +x_11 = l_Lean_Meta_mkDiagSummary___closed__2; +x_12 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1(x_1, x_10, x_2, x_11); +x_13 = l_Array_isEmpty___rarg(x_12); +if (x_13 == 0) +{ +lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_array_get_size(x_12); +x_15 = lean_usize_of_nat(x_14); +x_16 = 0; +x_17 = l_Lean_Meta_collectAboveThreshold___rarg___closed__1; +x_18 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8(x_12, x_15, x_16, x_17, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_18) == 0) +{ +uint8_t x_19; +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_lt(x_21, x_14); +lean_dec(x_14); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_12); +x_23 = l_Lean_Meta_mkDiagSummary___closed__3; +x_24 = l___private_Init_GetElem_0__outOfBounds___rarg(x_23); +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_20); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_18, 0, x_26); +return x_18; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_array_fget(x_12, x_21); +lean_dec(x_12); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(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); +lean_ctor_set(x_18, 0, x_29); +return x_18; +} +} +else +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_30 = lean_ctor_get(x_18, 0); +x_31 = lean_ctor_get(x_18, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_18); +x_32 = lean_unsigned_to_nat(0u); +x_33 = lean_nat_dec_lt(x_32, x_14); +lean_dec(x_14); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_12); +x_34 = l_Lean_Meta_mkDiagSummary___closed__3; +x_35 = l___private_Init_GetElem_0__outOfBounds___rarg(x_34); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_31); +return x_38; +} +else +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_array_fget(x_12, x_32); +lean_dec(x_12); +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_30); +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_31); +return x_42; +} +} +} +else +{ +uint8_t x_43; +lean_dec(x_14); +lean_dec(x_12); +x_43 = !lean_is_exclusive(x_18); +if (x_43 == 0) +{ +return x_18; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_18, 0); +x_45 = lean_ctor_get(x_18, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_18); +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 +{ +lean_object* x_47; lean_object* x_48; +lean_dec(x_12); +x_47 = l_Lean_Meta_instInhabitedDiagSummary___closed__1; +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_7); +return x_48; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___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; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_mkDiagSummary___spec__5___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___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_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_mkDiagSummary___spec__6___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Array_qsort_sort___at_Lean_Meta_mkDiagSummary___spec__7(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_mkDiagSummary___spec__1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +size_t x_10; size_t x_11; lean_object* x_12; +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_Meta_mkDiagSummary___spec__8(x_1, x_10, x_11, 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_1); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummary___boxed(lean_object* x_1, 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_Meta_mkDiagSummary(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 uint8_t l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1(lean_object* x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; lean_object* x_5; +lean_inc(x_3); +lean_inc(x_1); +x_4 = lean_get_reducibility_status(x_1, x_3); +x_5 = lean_box(x_4); +if (lean_obj_tag(x_5) == 1) +{ +uint8_t x_6; +x_6 = l_Lean_Meta_isInstanceCore(x_1, x_3); +lean_dec(x_1); +if (x_6 == 0) +{ +if (x_2 == 0) +{ +uint8_t x_7; +x_7 = 1; +return x_7; +} +else +{ +uint8_t x_8; +x_8 = 0; +return x_8; +} +} +else +{ +return x_2; +} +} +else +{ +uint8_t x_9; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +x_9 = 0; +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded(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; lean_object* x_14; +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); +x_12 = lean_box(x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1___boxed), 3, 2); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_12); +x_14 = l_Lean_Meta_mkDiagSummary(x_1, x_13, x_3, x_4, x_5, x_6, x_10); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_2); +lean_dec(x_2); +x_5 = l_Lean_Meta_mkDiagSummaryForUnfolded___lambda__1(x_1, x_4, x_3); +x_6 = lean_box(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfolded___boxed(lean_object* x_1, 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_mkDiagSummaryForUnfolded(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_get_reducibility_status(x_1, x_2); +x_4 = lean_box(x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = 1; +return x_5; +} +else +{ +uint8_t x_6; +lean_dec(x_4); +x_6 = 0; +return x_6; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_7 = lean_st_ref_get(x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1___boxed), 2, 1); +lean_closure_set(x_11, 0, x_10); +x_12 = l_Lean_Meta_mkDiagSummary(x_1, x_11, x_2, x_3, x_4, x_5, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___lambda__1(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUnfoldedReducible___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_Meta_mkDiagSummaryForUnfoldedReducible(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 uint8_t l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_st_ref_get(x_2, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_ctor_get(x_7, 4); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_9, 2); +lean_inc(x_10); +lean_dec(x_9); +x_11 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1; +x_12 = l_Lean_Meta_mkDiagSummary(x_10, x_11, x_1, x_2, x_3, x_4, x_8); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_mkDiagSummaryForUsedInstances___lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_mkDiagSummaryForUsedInstances___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_Meta_mkDiagSummaryForUsedInstances(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_6; +} +} +static lean_object* _init_l_Lean_Meta_appendSection___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" (max: ", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_appendSection___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(", num: ", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_appendSection___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("):", 2); +return x_1; +} +} +static double _init_l_Lean_Meta_appendSection___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_appendSection(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = l_Lean_Meta_DiagSummary_isEmpty(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; 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; double 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; +x_6 = l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1; +x_7 = lean_string_append(x_6, x_3); +x_8 = l_Lean_Meta_appendSection___closed__1; +x_9 = lean_string_append(x_7, x_8); +x_10 = lean_ctor_get(x_4, 1); +lean_inc(x_10); +x_11 = l___private_Init_Data_Repr_0__Nat_reprFast(x_10); +x_12 = lean_string_append(x_9, x_11); +lean_dec(x_11); +x_13 = l_Lean_Meta_appendSection___closed__2; +x_14 = lean_string_append(x_12, x_13); +x_15 = lean_ctor_get(x_4, 0); +lean_inc(x_15); +lean_dec(x_4); +x_16 = lean_array_get_size(x_15); +x_17 = l___private_Init_Data_Repr_0__Nat_reprFast(x_16); +x_18 = lean_string_append(x_14, x_17); +lean_dec(x_17); +x_19 = l_Lean_Meta_appendSection___closed__3; +x_20 = lean_string_append(x_18, x_19); +x_21 = l_Lean_Meta_appendSection___closed__4; +x_22 = 1; +x_23 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_23, 0, x_2); +lean_ctor_set(x_23, 1, x_6); +lean_ctor_set_float(x_23, sizeof(void*)*2, x_21); +lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_21); +lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_22); +x_24 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_24, 0, x_20); +x_25 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_26, 0, x_23); +lean_ctor_set(x_26, 1, x_25); +lean_ctor_set(x_26, 2, x_15); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_1); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +else +{ +lean_dec(x_4); +lean_dec(x_2); +return x_1; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_appendSection___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_appendSection(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagExt; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("reduction", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___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_reportDiag___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unfolded declarations", 21); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unfolded instances", 18); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unfolded reducible declarations", 31); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__7() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("type_class", 10); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_reportDiag___closed__7; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("used instances", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__10() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("def_eq", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_reportDiag___closed__10; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__12() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("heuristic for solving `f a =\?= f b`", 35); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__13() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("kernel", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__14() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_reportDiag___closed__13; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("use `set_option diagnostics.threshold ` to control threshold for reporting counters", 88); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_reportDiag___closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_reportDiag___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_Lean_Meta_reportDiag___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_reportDiag___closed__16; +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_Meta_reportDiag(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; +x_6 = l_Lean_isDiagnosticsEnabled(x_3, x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_unbox(x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +uint8_t x_9; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_9 = !lean_is_exclusive(x_6); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_6, 0); +lean_dec(x_10); +x_11 = lean_box(0); +lean_ctor_set(x_6, 0, x_11); +return x_6; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_6, 1); +lean_inc(x_12); +lean_dec(x_6); +x_13 = lean_box(0); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_15 = lean_ctor_get(x_6, 1); +lean_inc(x_15); +lean_dec(x_6); +x_16 = lean_st_ref_get(x_2, 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_ctor_get(x_17, 4); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +lean_dec(x_19); +x_21 = 0; +lean_inc(x_20); +x_22 = l_Lean_Meta_mkDiagSummaryForUnfolded(x_20, x_21, x_1, x_2, x_3, x_4, x_18); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = 1; +lean_inc(x_20); +x_26 = l_Lean_Meta_mkDiagSummaryForUnfolded(x_20, x_25, x_1, x_2, x_3, x_4, x_24); +if (lean_obj_tag(x_26) == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = l_Lean_Meta_mkDiagSummaryForUnfoldedReducible(x_20, x_1, x_2, x_3, x_4, x_28); +if (lean_obj_tag(x_29) == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; 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); +x_32 = lean_st_ref_get(x_2, x_31); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_ctor_get(x_33, 4); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1; +x_38 = l_Lean_Meta_mkDiagSummary(x_36, x_37, x_1, x_2, x_3, x_4, x_34); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_38, 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_Meta_mkDiagSummaryForUsedInstances(x_1, x_2, x_3, x_4, x_40); +if (lean_obj_tag(x_41) == 0) +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; 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_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_st_ref_get(x_4, x_43); +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +lean_dec(x_45); +x_48 = l_Lean_Kernel_instInhabitedDiagnostics; +x_49 = l_Lean_Meta_reportDiag___closed__1; +x_50 = l_Lean_EnvExtension_getState___rarg(x_48, x_49, x_47); +lean_dec(x_47); +x_51 = lean_ctor_get(x_50, 0); +lean_inc(x_51); +lean_dec(x_50); +x_52 = l_Lean_Meta_mkDiagSummary(x_51, x_37, x_1, x_2, x_3, x_4, x_46); +if (lean_obj_tag(x_52) == 0) +{ +uint8_t x_53; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = lean_ctor_get(x_52, 0); +x_55 = lean_ctor_get(x_52, 1); +x_56 = l_Lean_Meta_DiagSummary_isEmpty(x_23); +if (x_56 == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; +lean_free_object(x_52); +x_57 = l_Lean_MessageData_nil; +x_58 = l_Lean_Meta_reportDiag___closed__3; +x_59 = l_Lean_Meta_reportDiag___closed__4; +x_60 = l_Lean_Meta_appendSection(x_57, x_58, x_59, x_23); +x_61 = l_Lean_Meta_reportDiag___closed__5; +x_62 = l_Lean_Meta_appendSection(x_60, x_58, x_61, x_27); +x_63 = l_Lean_Meta_reportDiag___closed__6; +x_64 = l_Lean_Meta_appendSection(x_62, x_58, x_63, x_30); +x_65 = l_Lean_Meta_reportDiag___closed__8; +x_66 = l_Lean_Meta_reportDiag___closed__9; +x_67 = l_Lean_Meta_appendSection(x_64, x_65, x_66, x_42); +x_68 = l_Lean_Meta_reportDiag___closed__11; +x_69 = l_Lean_Meta_reportDiag___closed__12; +x_70 = l_Lean_Meta_appendSection(x_67, x_68, x_69, x_39); +x_71 = l_Lean_Meta_reportDiag___closed__14; +x_72 = l_Lean_Meta_appendSection(x_70, x_71, x_59, x_54); +x_73 = l_Lean_Meta_reportDiag___closed__17; +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +x_75 = 0; +x_76 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_74, x_75, x_1, x_2, x_3, x_4, x_55); +return x_76; +} +else +{ +uint8_t x_77; +x_77 = l_Lean_Meta_DiagSummary_isEmpty(x_27); +if (x_77 == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; 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_free_object(x_52); +x_78 = l_Lean_MessageData_nil; +x_79 = l_Lean_Meta_reportDiag___closed__3; +x_80 = l_Lean_Meta_reportDiag___closed__4; +x_81 = l_Lean_Meta_appendSection(x_78, x_79, x_80, x_23); +x_82 = l_Lean_Meta_reportDiag___closed__5; +x_83 = l_Lean_Meta_appendSection(x_81, x_79, x_82, x_27); +x_84 = l_Lean_Meta_reportDiag___closed__6; +x_85 = l_Lean_Meta_appendSection(x_83, x_79, x_84, x_30); +x_86 = l_Lean_Meta_reportDiag___closed__8; +x_87 = l_Lean_Meta_reportDiag___closed__9; +x_88 = l_Lean_Meta_appendSection(x_85, x_86, x_87, x_42); +x_89 = l_Lean_Meta_reportDiag___closed__11; +x_90 = l_Lean_Meta_reportDiag___closed__12; +x_91 = l_Lean_Meta_appendSection(x_88, x_89, x_90, x_39); +x_92 = l_Lean_Meta_reportDiag___closed__14; +x_93 = l_Lean_Meta_appendSection(x_91, x_92, x_80, x_54); +x_94 = l_Lean_Meta_reportDiag___closed__17; +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 = 0; +x_97 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_95, x_96, x_1, x_2, x_3, x_4, x_55); +return x_97; +} +else +{ +uint8_t x_98; +x_98 = l_Lean_Meta_DiagSummary_isEmpty(x_30); +if (x_98 == 0) +{ +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_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; uint8_t x_117; lean_object* x_118; +lean_free_object(x_52); +x_99 = l_Lean_MessageData_nil; +x_100 = l_Lean_Meta_reportDiag___closed__3; +x_101 = l_Lean_Meta_reportDiag___closed__4; +x_102 = l_Lean_Meta_appendSection(x_99, x_100, x_101, x_23); +x_103 = l_Lean_Meta_reportDiag___closed__5; +x_104 = l_Lean_Meta_appendSection(x_102, x_100, x_103, x_27); +x_105 = l_Lean_Meta_reportDiag___closed__6; +x_106 = l_Lean_Meta_appendSection(x_104, x_100, x_105, x_30); +x_107 = l_Lean_Meta_reportDiag___closed__8; +x_108 = l_Lean_Meta_reportDiag___closed__9; +x_109 = l_Lean_Meta_appendSection(x_106, x_107, x_108, x_42); +x_110 = l_Lean_Meta_reportDiag___closed__11; +x_111 = l_Lean_Meta_reportDiag___closed__12; +x_112 = l_Lean_Meta_appendSection(x_109, x_110, x_111, x_39); +x_113 = l_Lean_Meta_reportDiag___closed__14; +x_114 = l_Lean_Meta_appendSection(x_112, x_113, x_101, x_54); +x_115 = l_Lean_Meta_reportDiag___closed__17; +x_116 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_116, 0, x_114); +lean_ctor_set(x_116, 1, x_115); +x_117 = 0; +x_118 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_116, x_117, x_1, x_2, x_3, x_4, x_55); +return x_118; +} +else +{ +uint8_t x_119; +x_119 = l_Lean_Meta_DiagSummary_isEmpty(x_39); +if (x_119 == 0) +{ +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; uint8_t x_138; lean_object* x_139; +lean_free_object(x_52); +x_120 = l_Lean_MessageData_nil; +x_121 = l_Lean_Meta_reportDiag___closed__3; +x_122 = l_Lean_Meta_reportDiag___closed__4; +x_123 = l_Lean_Meta_appendSection(x_120, x_121, x_122, x_23); +x_124 = l_Lean_Meta_reportDiag___closed__5; +x_125 = l_Lean_Meta_appendSection(x_123, x_121, x_124, x_27); +x_126 = l_Lean_Meta_reportDiag___closed__6; +x_127 = l_Lean_Meta_appendSection(x_125, x_121, x_126, x_30); +x_128 = l_Lean_Meta_reportDiag___closed__8; +x_129 = l_Lean_Meta_reportDiag___closed__9; +x_130 = l_Lean_Meta_appendSection(x_127, x_128, x_129, x_42); +x_131 = l_Lean_Meta_reportDiag___closed__11; +x_132 = l_Lean_Meta_reportDiag___closed__12; +x_133 = l_Lean_Meta_appendSection(x_130, x_131, x_132, x_39); +x_134 = l_Lean_Meta_reportDiag___closed__14; +x_135 = l_Lean_Meta_appendSection(x_133, x_134, x_122, x_54); +x_136 = l_Lean_Meta_reportDiag___closed__17; +x_137 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_137, 0, x_135); +lean_ctor_set(x_137, 1, x_136); +x_138 = 0; +x_139 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_137, x_138, x_1, x_2, x_3, x_4, x_55); +return x_139; +} +else +{ +uint8_t x_140; +x_140 = l_Lean_Meta_DiagSummary_isEmpty(x_42); +if (x_140 == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; uint8_t x_159; lean_object* x_160; +lean_free_object(x_52); +x_141 = l_Lean_MessageData_nil; +x_142 = l_Lean_Meta_reportDiag___closed__3; +x_143 = l_Lean_Meta_reportDiag___closed__4; +x_144 = l_Lean_Meta_appendSection(x_141, x_142, x_143, x_23); +x_145 = l_Lean_Meta_reportDiag___closed__5; +x_146 = l_Lean_Meta_appendSection(x_144, x_142, x_145, x_27); +x_147 = l_Lean_Meta_reportDiag___closed__6; +x_148 = l_Lean_Meta_appendSection(x_146, x_142, x_147, x_30); +x_149 = l_Lean_Meta_reportDiag___closed__8; +x_150 = l_Lean_Meta_reportDiag___closed__9; +x_151 = l_Lean_Meta_appendSection(x_148, x_149, x_150, x_42); +x_152 = l_Lean_Meta_reportDiag___closed__11; +x_153 = l_Lean_Meta_reportDiag___closed__12; +x_154 = l_Lean_Meta_appendSection(x_151, x_152, x_153, x_39); +x_155 = l_Lean_Meta_reportDiag___closed__14; +x_156 = l_Lean_Meta_appendSection(x_154, x_155, x_143, x_54); +x_157 = l_Lean_Meta_reportDiag___closed__17; +x_158 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_158, 0, x_156); +lean_ctor_set(x_158, 1, x_157); +x_159 = 0; +x_160 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_158, x_159, x_1, x_2, x_3, x_4, x_55); +return x_160; +} +else +{ +lean_object* x_161; +lean_dec(x_54); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_161 = lean_box(0); +lean_ctor_set(x_52, 0, x_161); +return x_52; +} +} +} +} +} +} +else +{ +lean_object* x_162; lean_object* x_163; uint8_t x_164; +x_162 = lean_ctor_get(x_52, 0); +x_163 = lean_ctor_get(x_52, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_52); +x_164 = l_Lean_Meta_DiagSummary_isEmpty(x_23); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_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; uint8_t x_183; lean_object* x_184; +x_165 = l_Lean_MessageData_nil; +x_166 = l_Lean_Meta_reportDiag___closed__3; +x_167 = l_Lean_Meta_reportDiag___closed__4; +x_168 = l_Lean_Meta_appendSection(x_165, x_166, x_167, x_23); +x_169 = l_Lean_Meta_reportDiag___closed__5; +x_170 = l_Lean_Meta_appendSection(x_168, x_166, x_169, x_27); +x_171 = l_Lean_Meta_reportDiag___closed__6; +x_172 = l_Lean_Meta_appendSection(x_170, x_166, x_171, x_30); +x_173 = l_Lean_Meta_reportDiag___closed__8; +x_174 = l_Lean_Meta_reportDiag___closed__9; +x_175 = l_Lean_Meta_appendSection(x_172, x_173, x_174, x_42); +x_176 = l_Lean_Meta_reportDiag___closed__11; +x_177 = l_Lean_Meta_reportDiag___closed__12; +x_178 = l_Lean_Meta_appendSection(x_175, x_176, x_177, x_39); +x_179 = l_Lean_Meta_reportDiag___closed__14; +x_180 = l_Lean_Meta_appendSection(x_178, x_179, x_167, x_162); +x_181 = l_Lean_Meta_reportDiag___closed__17; +x_182 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_182, 0, x_180); +lean_ctor_set(x_182, 1, x_181); +x_183 = 0; +x_184 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_182, x_183, x_1, x_2, x_3, x_4, x_163); +return x_184; +} +else +{ +uint8_t x_185; +x_185 = l_Lean_Meta_DiagSummary_isEmpty(x_27); +if (x_185 == 0) +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; uint8_t x_204; lean_object* x_205; +x_186 = l_Lean_MessageData_nil; +x_187 = l_Lean_Meta_reportDiag___closed__3; +x_188 = l_Lean_Meta_reportDiag___closed__4; +x_189 = l_Lean_Meta_appendSection(x_186, x_187, x_188, x_23); +x_190 = l_Lean_Meta_reportDiag___closed__5; +x_191 = l_Lean_Meta_appendSection(x_189, x_187, x_190, x_27); +x_192 = l_Lean_Meta_reportDiag___closed__6; +x_193 = l_Lean_Meta_appendSection(x_191, x_187, x_192, x_30); +x_194 = l_Lean_Meta_reportDiag___closed__8; +x_195 = l_Lean_Meta_reportDiag___closed__9; +x_196 = l_Lean_Meta_appendSection(x_193, x_194, x_195, x_42); +x_197 = l_Lean_Meta_reportDiag___closed__11; +x_198 = l_Lean_Meta_reportDiag___closed__12; +x_199 = l_Lean_Meta_appendSection(x_196, x_197, x_198, x_39); +x_200 = l_Lean_Meta_reportDiag___closed__14; +x_201 = l_Lean_Meta_appendSection(x_199, x_200, x_188, x_162); +x_202 = l_Lean_Meta_reportDiag___closed__17; +x_203 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_203, 0, x_201); +lean_ctor_set(x_203, 1, x_202); +x_204 = 0; +x_205 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_203, x_204, x_1, x_2, x_3, x_4, x_163); +return x_205; +} +else +{ +uint8_t x_206; +x_206 = l_Lean_Meta_DiagSummary_isEmpty(x_30); +if (x_206 == 0) +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; lean_object* x_226; +x_207 = l_Lean_MessageData_nil; +x_208 = l_Lean_Meta_reportDiag___closed__3; +x_209 = l_Lean_Meta_reportDiag___closed__4; +x_210 = l_Lean_Meta_appendSection(x_207, x_208, x_209, x_23); +x_211 = l_Lean_Meta_reportDiag___closed__5; +x_212 = l_Lean_Meta_appendSection(x_210, x_208, x_211, x_27); +x_213 = l_Lean_Meta_reportDiag___closed__6; +x_214 = l_Lean_Meta_appendSection(x_212, x_208, x_213, x_30); +x_215 = l_Lean_Meta_reportDiag___closed__8; +x_216 = l_Lean_Meta_reportDiag___closed__9; +x_217 = l_Lean_Meta_appendSection(x_214, x_215, x_216, x_42); +x_218 = l_Lean_Meta_reportDiag___closed__11; +x_219 = l_Lean_Meta_reportDiag___closed__12; +x_220 = l_Lean_Meta_appendSection(x_217, x_218, x_219, x_39); +x_221 = l_Lean_Meta_reportDiag___closed__14; +x_222 = l_Lean_Meta_appendSection(x_220, x_221, x_209, x_162); +x_223 = l_Lean_Meta_reportDiag___closed__17; +x_224 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_224, 0, x_222); +lean_ctor_set(x_224, 1, x_223); +x_225 = 0; +x_226 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_224, x_225, x_1, x_2, x_3, x_4, x_163); +return x_226; +} +else +{ +uint8_t x_227; +x_227 = l_Lean_Meta_DiagSummary_isEmpty(x_39); +if (x_227 == 0) +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; lean_object* x_247; +x_228 = l_Lean_MessageData_nil; +x_229 = l_Lean_Meta_reportDiag___closed__3; +x_230 = l_Lean_Meta_reportDiag___closed__4; +x_231 = l_Lean_Meta_appendSection(x_228, x_229, x_230, x_23); +x_232 = l_Lean_Meta_reportDiag___closed__5; +x_233 = l_Lean_Meta_appendSection(x_231, x_229, x_232, x_27); +x_234 = l_Lean_Meta_reportDiag___closed__6; +x_235 = l_Lean_Meta_appendSection(x_233, x_229, x_234, x_30); +x_236 = l_Lean_Meta_reportDiag___closed__8; +x_237 = l_Lean_Meta_reportDiag___closed__9; +x_238 = l_Lean_Meta_appendSection(x_235, x_236, x_237, x_42); +x_239 = l_Lean_Meta_reportDiag___closed__11; +x_240 = l_Lean_Meta_reportDiag___closed__12; +x_241 = l_Lean_Meta_appendSection(x_238, x_239, x_240, x_39); +x_242 = l_Lean_Meta_reportDiag___closed__14; +x_243 = l_Lean_Meta_appendSection(x_241, x_242, x_230, x_162); +x_244 = l_Lean_Meta_reportDiag___closed__17; +x_245 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_245, 0, x_243); +lean_ctor_set(x_245, 1, x_244); +x_246 = 0; +x_247 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_245, x_246, x_1, x_2, x_3, x_4, x_163); +return x_247; +} +else +{ +uint8_t x_248; +x_248 = l_Lean_Meta_DiagSummary_isEmpty(x_42); +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; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t x_267; lean_object* x_268; +x_249 = l_Lean_MessageData_nil; +x_250 = l_Lean_Meta_reportDiag___closed__3; +x_251 = l_Lean_Meta_reportDiag___closed__4; +x_252 = l_Lean_Meta_appendSection(x_249, x_250, x_251, x_23); +x_253 = l_Lean_Meta_reportDiag___closed__5; +x_254 = l_Lean_Meta_appendSection(x_252, x_250, x_253, x_27); +x_255 = l_Lean_Meta_reportDiag___closed__6; +x_256 = l_Lean_Meta_appendSection(x_254, x_250, x_255, x_30); +x_257 = l_Lean_Meta_reportDiag___closed__8; +x_258 = l_Lean_Meta_reportDiag___closed__9; +x_259 = l_Lean_Meta_appendSection(x_256, x_257, x_258, x_42); +x_260 = l_Lean_Meta_reportDiag___closed__11; +x_261 = l_Lean_Meta_reportDiag___closed__12; +x_262 = l_Lean_Meta_appendSection(x_259, x_260, x_261, x_39); +x_263 = l_Lean_Meta_reportDiag___closed__14; +x_264 = l_Lean_Meta_appendSection(x_262, x_263, x_251, x_162); +x_265 = l_Lean_Meta_reportDiag___closed__17; +x_266 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_266, 0, x_264); +lean_ctor_set(x_266, 1, x_265); +x_267 = 0; +x_268 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_266, x_267, x_1, x_2, x_3, x_4, x_163); +return x_268; +} +else +{ +lean_object* x_269; lean_object* x_270; +lean_dec(x_162); +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_269 = lean_box(0); +x_270 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_163); +return x_270; +} +} +} +} +} +} +} +else +{ +uint8_t x_271; +lean_dec(x_42); +lean_dec(x_39); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_271 = !lean_is_exclusive(x_52); +if (x_271 == 0) +{ +return x_52; +} +else +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_272 = lean_ctor_get(x_52, 0); +x_273 = lean_ctor_get(x_52, 1); +lean_inc(x_273); +lean_inc(x_272); +lean_dec(x_52); +x_274 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_274, 0, x_272); +lean_ctor_set(x_274, 1, x_273); +return x_274; +} +} +} +else +{ +uint8_t x_275; +lean_dec(x_39); +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_275 = !lean_is_exclusive(x_41); +if (x_275 == 0) +{ +return x_41; +} +else +{ +lean_object* x_276; lean_object* x_277; lean_object* x_278; +x_276 = lean_ctor_get(x_41, 0); +x_277 = lean_ctor_get(x_41, 1); +lean_inc(x_277); +lean_inc(x_276); +lean_dec(x_41); +x_278 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_278, 0, x_276); +lean_ctor_set(x_278, 1, x_277); +return x_278; +} +} +} +else +{ +uint8_t x_279; +lean_dec(x_30); +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_279 = !lean_is_exclusive(x_38); +if (x_279 == 0) +{ +return x_38; +} +else +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; +x_280 = lean_ctor_get(x_38, 0); +x_281 = lean_ctor_get(x_38, 1); +lean_inc(x_281); +lean_inc(x_280); +lean_dec(x_38); +x_282 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_282, 0, x_280); +lean_ctor_set(x_282, 1, x_281); +return x_282; +} +} +} +else +{ +uint8_t x_283; +lean_dec(x_27); +lean_dec(x_23); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_283 = !lean_is_exclusive(x_29); +if (x_283 == 0) +{ +return x_29; +} +else +{ +lean_object* x_284; lean_object* x_285; lean_object* x_286; +x_284 = lean_ctor_get(x_29, 0); +x_285 = lean_ctor_get(x_29, 1); +lean_inc(x_285); +lean_inc(x_284); +lean_dec(x_29); +x_286 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_286, 0, x_284); +lean_ctor_set(x_286, 1, x_285); +return x_286; +} +} +} +else +{ +uint8_t x_287; +lean_dec(x_23); +lean_dec(x_20); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_287 = !lean_is_exclusive(x_26); +if (x_287 == 0) +{ +return x_26; +} +else +{ +lean_object* x_288; lean_object* x_289; lean_object* x_290; +x_288 = lean_ctor_get(x_26, 0); +x_289 = lean_ctor_get(x_26, 1); +lean_inc(x_289); +lean_inc(x_288); +lean_dec(x_26); +x_290 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_290, 0, x_288); +lean_ctor_set(x_290, 1, x_289); +return x_290; +} +} +} +else +{ +uint8_t x_291; +lean_dec(x_20); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_291 = !lean_is_exclusive(x_22); +if (x_291 == 0) +{ +return x_22; +} +else +{ +lean_object* x_292; lean_object* x_293; lean_object* x_294; +x_292 = lean_ctor_get(x_22, 0); +x_293 = lean_ctor_get(x_22, 1); +lean_inc(x_293); +lean_inc(x_292); +lean_dec(x_22); +x_294 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); +return x_294; +} +} +} +} +} +lean_object* initialize_Lean_PrettyPrinter(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Instances(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Diagnostics(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_PrettyPrinter(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_Instances(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_collectAboveThreshold___rarg___closed__1 = _init_l_Lean_Meta_collectAboveThreshold___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_collectAboveThreshold___rarg___closed__1); +l_Lean_Meta_subCounters___rarg___closed__1 = _init_l_Lean_Meta_subCounters___rarg___closed__1(); +lean_mark_persistent(l_Lean_Meta_subCounters___rarg___closed__1); +l_Lean_Meta_subCounters___rarg___closed__2 = _init_l_Lean_Meta_subCounters___rarg___closed__2(); +lean_mark_persistent(l_Lean_Meta_subCounters___rarg___closed__2); +l_Lean_Meta_DiagSummary_data___default = _init_l_Lean_Meta_DiagSummary_data___default(); +lean_mark_persistent(l_Lean_Meta_DiagSummary_data___default); +l_Lean_Meta_DiagSummary_max___default = _init_l_Lean_Meta_DiagSummary_max___default(); +lean_mark_persistent(l_Lean_Meta_DiagSummary_max___default); +l_Lean_Meta_instInhabitedDiagSummary___closed__1 = _init_l_Lean_Meta_instInhabitedDiagSummary___closed__1(); +lean_mark_persistent(l_Lean_Meta_instInhabitedDiagSummary___closed__1); +l_Lean_Meta_instInhabitedDiagSummary = _init_l_Lean_Meta_instInhabitedDiagSummary(); +lean_mark_persistent(l_Lean_Meta_instInhabitedDiagSummary); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__7); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__8); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__9); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__10); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__11); +l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__12 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_mkDiagSummary___spec__8___closed__12); +l_Lean_Meta_mkDiagSummary___closed__1 = _init_l_Lean_Meta_mkDiagSummary___closed__1(); +lean_mark_persistent(l_Lean_Meta_mkDiagSummary___closed__1); +l_Lean_Meta_mkDiagSummary___closed__2 = _init_l_Lean_Meta_mkDiagSummary___closed__2(); +lean_mark_persistent(l_Lean_Meta_mkDiagSummary___closed__2); +l_Lean_Meta_mkDiagSummary___closed__3 = _init_l_Lean_Meta_mkDiagSummary___closed__3(); +lean_mark_persistent(l_Lean_Meta_mkDiagSummary___closed__3); +l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1 = _init_l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1(); +lean_mark_persistent(l_Lean_Meta_mkDiagSummaryForUsedInstances___closed__1); +l_Lean_Meta_appendSection___closed__1 = _init_l_Lean_Meta_appendSection___closed__1(); +lean_mark_persistent(l_Lean_Meta_appendSection___closed__1); +l_Lean_Meta_appendSection___closed__2 = _init_l_Lean_Meta_appendSection___closed__2(); +lean_mark_persistent(l_Lean_Meta_appendSection___closed__2); +l_Lean_Meta_appendSection___closed__3 = _init_l_Lean_Meta_appendSection___closed__3(); +lean_mark_persistent(l_Lean_Meta_appendSection___closed__3); +l_Lean_Meta_appendSection___closed__4 = _init_l_Lean_Meta_appendSection___closed__4(); +l_Lean_Meta_reportDiag___closed__1 = _init_l_Lean_Meta_reportDiag___closed__1(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__1); +l_Lean_Meta_reportDiag___closed__2 = _init_l_Lean_Meta_reportDiag___closed__2(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__2); +l_Lean_Meta_reportDiag___closed__3 = _init_l_Lean_Meta_reportDiag___closed__3(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__3); +l_Lean_Meta_reportDiag___closed__4 = _init_l_Lean_Meta_reportDiag___closed__4(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__4); +l_Lean_Meta_reportDiag___closed__5 = _init_l_Lean_Meta_reportDiag___closed__5(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__5); +l_Lean_Meta_reportDiag___closed__6 = _init_l_Lean_Meta_reportDiag___closed__6(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__6); +l_Lean_Meta_reportDiag___closed__7 = _init_l_Lean_Meta_reportDiag___closed__7(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__7); +l_Lean_Meta_reportDiag___closed__8 = _init_l_Lean_Meta_reportDiag___closed__8(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__8); +l_Lean_Meta_reportDiag___closed__9 = _init_l_Lean_Meta_reportDiag___closed__9(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__9); +l_Lean_Meta_reportDiag___closed__10 = _init_l_Lean_Meta_reportDiag___closed__10(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__10); +l_Lean_Meta_reportDiag___closed__11 = _init_l_Lean_Meta_reportDiag___closed__11(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__11); +l_Lean_Meta_reportDiag___closed__12 = _init_l_Lean_Meta_reportDiag___closed__12(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__12); +l_Lean_Meta_reportDiag___closed__13 = _init_l_Lean_Meta_reportDiag___closed__13(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__13); +l_Lean_Meta_reportDiag___closed__14 = _init_l_Lean_Meta_reportDiag___closed__14(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__14); +l_Lean_Meta_reportDiag___closed__15 = _init_l_Lean_Meta_reportDiag___closed__15(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__15); +l_Lean_Meta_reportDiag___closed__16 = _init_l_Lean_Meta_reportDiag___closed__16(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__16); +l_Lean_Meta_reportDiag___closed__17 = _init_l_Lean_Meta_reportDiag___closed__17(); +lean_mark_persistent(l_Lean_Meta_reportDiag___closed__17); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Eqns.c b/stage0/stdlib/Lean/Meta/Eqns.c index 740c4953184d..1cb2a8cf63cb 100644 --- a/stage0/stdlib/Lean/Meta/Eqns.c +++ b/stage0/stdlib/Lean/Meta/Eqns.c @@ -68,7 +68,6 @@ lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_EnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_registerGetUnfoldEqnFn(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_isUnfoldReservedNameSuffix(lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Meta_getEqnsFor_x3f___closed__1; @@ -122,7 +121,6 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2109____lambda__2___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_getEqnsFnsRef; lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Meta_Eqns_0__Lean_Meta_registerEqnThms___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -136,8 +134,8 @@ static lean_object* l_Lean_Meta_getUnfoldEqnFor_x3f___lambda__1___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isEqnThm_x3f(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Meta_getUnfoldEqnFor_x3f___spec__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Meta_getUnfoldEqnFor_x3f___spec__1___closed__5; static lean_object* l_Lean_Meta_markAsRecursive___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2109____lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -219,6 +217,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2109____lambd static lean_object* l_Lean_Meta_getEqnsFor_x3f___closed__2; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2109____lambda__2___closed__3; size_t lean_usize_add(size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_ensureReservedNameAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_493____closed__1; lean_object* lean_array_uget(lean_object*, size_t); @@ -249,6 +248,7 @@ lean_object* l_Lean_mkTagDeclarationExtension(lean_object*, lean_object*); static lean_object* l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_isEqnThm_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_unfoldThmSuffix___closed__1; +lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_2109____lambda__2___closed__20; static lean_object* l_List_forIn_loop___at_Lean_Meta_getUnfoldEqnFor_x3f___spec__1___closed__1; static lean_object* l_Lean_Meta_isEqnThm_x3f___closed__1; @@ -633,44 +633,6 @@ x_3 = lean_box(x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 5); -x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_inc(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); -lean_ctor_set_tag(x_6, 1); -lean_ctor_set(x_6, 0, x_9); -return x_6; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_6, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_6); -lean_inc(x_5); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; -} -} -} static lean_object* _init_l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2___closed__1() { _start: { @@ -743,9 +705,7 @@ x_13 = l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedName x_14 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_14, 0, x_12); lean_ctor_set(x_14, 1, x_13); -x_15 = l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3(x_14, x_3, x_4, x_5); -lean_dec(x_4); -lean_dec(x_3); +x_15 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_14, x_3, x_4, x_5); return x_15; } } @@ -771,8 +731,6 @@ if (x_12 == 0) { lean_object* x_13; lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); x_13 = lean_box(0); lean_ctor_set(x_7, 0, x_13); @@ -803,8 +761,6 @@ if (x_18 == 0) { lean_object* x_19; lean_object* x_20; lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_1); x_19 = lean_box(0); x_20 = lean_alloc_ctor(0, 2, 0); @@ -826,8 +782,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_ensureEqnReservedNamesAvailable(lean_object { lean_object* x_5; lean_object* x_6; x_5 = l_Lean_Meta_unfoldThmSuffix; -lean_inc(x_3); -lean_inc(x_2); lean_inc(x_1); x_6 = l_Lean_ensureReservedNameAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__1(x_1, x_5, x_2, x_3, x_4); if (lean_obj_tag(x_6) == 0) @@ -838,6 +792,8 @@ lean_inc(x_7); lean_dec(x_6); x_8 = l_Lean_Meta_eqn1ThmSuffix; x_9 = l_Lean_ensureReservedNameAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__1(x_1, x_8, x_2, x_3, x_7); +lean_dec(x_3); +lean_dec(x_2); return x_9; } else @@ -867,14 +823,24 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__3(x_1, x_2, x_3, x_4); +lean_object* x_6; +x_6 = l_Lean_throwReservedNameNotAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_5; +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_ensureReservedNameAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___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_6; +x_6 = l_Lean_ensureReservedNameAvailable___at_Lean_Meta_ensureEqnReservedNamesAvailable___spec__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_6; } } LEAN_EXPORT uint8_t l_Lean_Meta_initFn____x40_Lean_Meta_Eqns___hyg_151____lambda__1(lean_object* x_1, lean_object* x_2) { diff --git a/stage0/stdlib/Lean/Meta/ExprDefEq.c b/stage0/stdlib/Lean/Meta/ExprDefEq.c index 9157a279fc01..6cc3e12c7064 100644 --- a/stage0/stdlib/Lean/Meta/ExprDefEq.c +++ b/stage0/stdlib/Lean/Meta/ExprDefEq.c @@ -15,13 +15,12 @@ extern "C" { #endif LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkCacheKey(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Core_instMonadTraceCoreM; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__8; lean_object* l_Lean_addTrace___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_is_matcher(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -51,12 +50,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeurist LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_profiler; lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCacheKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_whenUndefDo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProjDelta(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_is_level_def_eq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_isAbstractedUnassignedMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -84,25 +85,26 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_CheckAssig 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_EXPORT lean_object* l_Lean_Meta_CheckAssignmentQuick_check_visit___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__50___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(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_EXPORT lean_object* l_Lean_occursCheck_visitMVar___at_Lean_Meta_checkAssignment___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____closed__3; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__2; lean_object* l_Lean_getDelayedMVarAssignment_x3f___at___private_Lean_Meta_WHNF_0__Lean_Meta_whnfDelayedAssigned_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); -static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_DefEqCache_update___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__55(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_reduceNat_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__9; static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__8; static lean_object* l_Lean_Meta_CheckAssignment_run___closed__1; +double lean_float_div(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__42___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_addLetDeps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -130,11 +132,9 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkM uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getDefEqCacheKind___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_defaultCase___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__3(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__7; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_Meta_checkAssignment___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); @@ -172,7 +172,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprD LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__66___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__4; static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__6; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____closed__7; uint8_t l_Array_contains___at___private_Lean_Class_0__Lean_checkOutParam___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2___closed__3; @@ -197,6 +196,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isAssignab LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_proj___override(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bvar___override(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__31(lean_object*, lean_object*); @@ -242,6 +242,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeurist LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj___closed__1; +lean_object* lean_io_get_num_heartbeats(lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__14; @@ -251,9 +252,11 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___c LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__31___boxed(lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2; lean_object* l_Lean_Expr_mvar___override(lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__33___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadLiftT(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_HashSetImp_moveEntries___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__12(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); @@ -267,18 +270,16 @@ static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__10; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__5; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2; uint8_t l_Lean_ReducibilityHints_compare(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq_packedInstanceOf_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__59___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__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_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqLeft___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___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_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__29(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__63(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*); @@ -305,7 +306,6 @@ LEAN_EXPORT uint8_t l_Lean_Meta_CheckAssignmentQuick_check(uint8_t, lean_object* uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__49(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqRight___closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_simpAssignmentArgAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__4(lean_object*, lean_object*, lean_object*); @@ -323,6 +323,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__1(uint8_t, lea lean_object* l_Lean_LocalContext_getFVar_x21(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_checkAssignment___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_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqApp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__3; static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__5; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__21___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_outOfScopeExceptionId; @@ -332,7 +333,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignmen size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_List_elem___at_Lean_catchInternalIds___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803_(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____closed__4; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__41___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__32(lean_object*, lean_object*, size_t, size_t); @@ -340,7 +340,6 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0_ lean_object* l_Lean_Meta_recordDefEqHeuristic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_run___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__12(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4; lean_object* l___private_Lean_Expr_0__Lean_mkAppRangeAux(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -376,6 +375,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_Persisten lean_object* l_Lean_Meta_shouldReduceReducibleOnly(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_is_expr_def_eq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__44___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_assignConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -398,21 +398,22 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFV lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom___spec__2(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__11; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkFVar___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_instantiateExprMVars___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___closed__18; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___spec__2(lean_object*, lean_object*, size_t, size_t); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); lean_object* l_Lean_Meta_SavedState_restore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__7(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__5; LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__11(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -427,7 +428,6 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_CheckAssignment lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__37(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__9(lean_object*, lean_object*, size_t, size_t); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__9; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_45____closed__1; lean_object* l_ReaderT_instMonadLift(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkAuxMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -461,6 +461,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDela lean_object* lean_array_pop(lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_etaEq___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___lambda__2___closed__1; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go_cont(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -468,6 +469,7 @@ uint64_t l___private_Lean_Expr_0__Lean_hashFVarId____x40_Lean_Expr___hyg_1674_(l LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__46___boxed(lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_withProofIrrelTransparency___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__8(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -478,7 +480,7 @@ uint8_t l_Lean_LocalContext_containsFVar(lean_object*, lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isHeadBetaTarget(lean_object*, uint8_t); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___boxed(lean_object**); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getDefEqCacheKind(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__60(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*); @@ -490,11 +492,13 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0_ LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickMVarMVar___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__2(lean_object*, lean_object*, size_t, size_t); +static double l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; extern lean_object* l_Lean_levelZero; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_cacheResult(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isDelayedAssigned___at_Lean_Meta_CheckAssignment_checkApp___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_expr_quick_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__26___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getTransparency(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox___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*); @@ -505,6 +509,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAss extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__25___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isNonTrivialRegular_isSimple(lean_object*, uint8_t); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4; lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_occursCheck___at_Lean_Meta_checkAssignment___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__2; @@ -521,7 +526,6 @@ static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10; lean_object* l_Lean_Expr_letE___override(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqBindingDomain(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -544,11 +548,11 @@ static lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___ra LEAN_EXPORT lean_object* l_Lean_Meta_checkAssignment___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go_cont___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfold___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_threshold; lean_object* l_Lean_MetavarContext_getExprAssignmentCore_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_CheckAssignment_checkApp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__60___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalContext_contains(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__8; lean_object* l_Lean_Meta_getResetPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__23___boxed(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___closed__4; @@ -576,14 +580,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVa LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__8(lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqUnitLike(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_tryHeuristic___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___rarg___closed__2; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_whnfCoreAtDefEq___closed__1; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__28___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__3; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -591,6 +595,7 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOth static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldComparingHeadsDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__7; lean_object* lean_usize_to_nat(size_t); @@ -598,7 +603,6 @@ size_t lean_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedProjectionFunctionInfo; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__5; 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; @@ -639,19 +643,17 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_DefEq LEAN_EXPORT uint8_t l_Lean_HashSetImp_contains___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDepsAux___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___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_Meta_ExprDefEq_0__Lean_Meta_isDefEqQuickOther___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__3; static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__30(lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_State_cache___default___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13; 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*); +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2; 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*, 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*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_CheckAssignment_checkMVar___spec__45(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_LocalContext_isSubPrefixOf(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isDefEqStringLit___closed__3; @@ -660,6 +662,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkM LEAN_EXPORT uint8_t l_List_elem___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_DefEqCache_update___spec__2(lean_object*, size_t, size_t, lean_object*, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Meta_CheckAssignment_checkMVar___spec__11(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_constLevels_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -674,13 +677,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_Persisten LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDepsAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkApp___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__9; lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_getCachedResult___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__13; +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isNonTrivialRegular___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_CheckAssignment_throwCheckAssignmentFailure___rarg___closed__2; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_CheckAssignment_checkApp___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__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*); @@ -708,17 +714,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unstuckMVa lean_object* l___private_Lean_Expr_0__Lean_Expr_etaExpandedAux(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar___rarg(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkFVar___spec__2(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_mkHashSet___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps___spec__1___boxed(lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____closed__2; -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_simpAssignmentArgAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__53___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignmentQuick_check_visit___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254_(lean_object*); lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_etaEq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -737,6 +742,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldDefE static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__3; lean_object* l_StateRefT_x27_lift(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); +static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1; lean_object* lean_local_ctx_erase(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_unfoldNonProjFnDefEq_packedInstanceOf_x3f___closed__1; LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_CheckAssignment_checkFVar___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -761,7 +767,6 @@ static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAs LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__33(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at_Lean_Meta_checkAssignment___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Meta_CheckAssignment_checkMVar___spec__4(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__11; LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Meta_CheckAssignment_checkMVar___spec__52(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashable; static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_CheckAssignment_checkFVar___spec__3___closed__1; @@ -780,7 +785,6 @@ LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_c LEAN_EXPORT lean_object* lean_is_expr_def_eq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isAssignable___boxed(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_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1; static lean_object* l_Lean_Meta_isDefEqStringLit___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_checkAssignment___spec__7(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj_isDefEqSingleton___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*); @@ -813,7 +817,7 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__57(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; +static double l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___boxed(lean_object**); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); lean_object* l_Lean_Meta_getFunInfoNArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -846,6 +850,7 @@ size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCache_update(lean_object*, uint8_t, lean_object*, uint8_t); static lean_object* l_Lean_Meta_CheckAssignment_checkMVar___closed__4; uint8_t lean_uint32_dec_lt(uint32_t, uint32_t); +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_assignToConstFun(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_DefEqCache_update___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__3(lean_object*, lean_object*); @@ -858,6 +863,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processCon LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__8; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__11; LEAN_EXPORT lean_object* l_Lean_occursCheck___at_Lean_Meta_checkAssignment___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_isDefEqStringLit___closed__2; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__2; @@ -895,6 +901,7 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__42(lean_object*, lean_object*, size_t, size_t); lean_object* l_Lean_Meta_useEtaStruct(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyMAux___at_Lean_Meta_CheckAssignment_checkMVar___spec__39___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12; static lean_object* l_Lean_Meta_CheckAssignment_instMonadCacheExprCheckAssignmentM___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgs___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_CheckAssignment_checkMVar___spec__41(lean_object*, lean_object*, size_t, size_t); @@ -925,8 +932,8 @@ extern lean_object* l_Lean_Meta_instAddMessageContextMetaM; lean_object* l_Lean_Meta_isDefEqOffset(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_List_elem___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_collectLetDeclsFrom_visit___spec__10___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7; lean_object* l_Lean_instMonadTraceOfMonadLift___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_CheckAssignment_checkMVar___spec__63___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMCtxMetaM; @@ -941,7 +948,6 @@ lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isExprDefEqExpensive___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_CheckAssignment_checkMVar___spec__30___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_isEtaUnassignedMVar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2; LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isNonTrivialRegular___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConst(lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); @@ -962,7 +968,6 @@ 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*); @@ -983,7 +988,7 @@ static lean_object* l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_run(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignment_process___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__13; +static lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1; lean_object* l_Lean_LocalContext_findFVar_x3f(lean_object*, lean_object*); uint8_t l_Lean_isStructureLike(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processAssignmentFOApprox_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -991,7 +996,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDef LEAN_EXPORT lean_object* l_Lean_Meta_DefEqCacheKind_noConfusion___rarg___lambda__1___boxed(lean_object*); lean_object* l_Lean_Meta_isProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getValue_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +static lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_throwOutOfScopeFVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_expandDelayedAssigned_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqLiteral____x40_Lean_Expr___hyg_31_(lean_object*, lean_object*); @@ -4466,7 +4471,7 @@ return x_99; } } } -static lean_object* _init_l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1() { +static lean_object* _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -4475,852 +4480,6 @@ x_2 = l_Lean_Expr_bvar___override(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, uint8_t x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_12 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1; -x_13 = l_Lean_Expr_app___override(x_2, x_12); -x_14 = l_Lean_Expr_lam___override(x_3, x_4, x_13, x_5); -x_20 = l_Lean_Meta_saveState___rarg(x_8, x_9, x_10, x_11); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_st_ref_take(x_8, x_22); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -x_26 = lean_ctor_get(x_23, 1); -lean_inc(x_26); -lean_dec(x_23); -x_27 = !lean_is_exclusive(x_24); -if (x_27 == 0) -{ -lean_object* x_28; uint8_t x_29; -x_28 = lean_ctor_get(x_24, 1); -lean_dec(x_28); -x_29 = !lean_is_exclusive(x_25); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_54; -x_30 = lean_ctor_get(x_25, 5); -lean_dec(x_30); -x_31 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__6; -lean_ctor_set(x_25, 5, x_31); -x_32 = lean_st_ref_set(x_8, x_24, x_26); -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); -x_34 = l_Lean_Meta_getResetPostponed(x_7, x_8, x_9, x_10, 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); -if (lean_is_exclusive(x_34)) { - lean_ctor_release(x_34, 0); - lean_ctor_release(x_34, 1); - x_37 = x_34; -} else { - lean_dec_ref(x_34); - x_37 = lean_box(0); -} -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_54 = lean_is_expr_def_eq(x_1, x_14, x_7, x_8, x_9, x_10, x_36); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; uint8_t x_56; -x_55 = lean_ctor_get(x_54, 0); -lean_inc(x_55); -x_56 = lean_unbox(x_55); -lean_dec(x_55); -if (x_56 == 0) -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -lean_dec(x_37); -lean_dec(x_35); -x_57 = lean_ctor_get(x_54, 1); -lean_inc(x_57); -lean_dec(x_54); -x_58 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_57); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_59 = lean_ctor_get(x_58, 1); -lean_inc(x_59); -lean_dec(x_58); -x_60 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_60; -x_16 = x_59; -goto block_19; -} -else -{ -lean_object* x_61; uint8_t x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_54, 1); -lean_inc(x_61); -lean_dec(x_54); -x_62 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_63 = l_Lean_Meta_processPostponed(x_6, x_62, x_7, x_8, x_9, x_10, x_61); -if (lean_obj_tag(x_63) == 0) -{ -lean_object* x_64; uint8_t x_65; -lean_dec(x_37); -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_unbox(x_64); -lean_dec(x_64); -if (x_65 == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -lean_dec(x_35); -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_66); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_68 = lean_ctor_get(x_67, 1); -lean_inc(x_68); -lean_dec(x_67); -x_69 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_69; -x_16 = x_68; -goto block_19; -} -else -{ -lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -lean_dec(x_21); -x_70 = lean_ctor_get(x_63, 1); -lean_inc(x_70); -lean_dec(x_63); -x_71 = l_Lean_Meta_getPostponed___rarg(x_8, x_9, x_10, x_70); -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_PersistentArray_append___rarg(x_35, x_72); -x_75 = l_Lean_Meta_setPostponed(x_74, x_7, x_8, x_9, x_10, x_73); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_76 = lean_ctor_get(x_75, 1); -lean_inc(x_76); -lean_dec(x_75); -x_77 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__8; -x_15 = x_77; -x_16 = x_76; -goto block_19; -} -} -else -{ -lean_object* x_78; lean_object* x_79; -lean_dec(x_35); -x_78 = lean_ctor_get(x_63, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_63, 1); -lean_inc(x_79); -lean_dec(x_63); -x_38 = x_78; -x_39 = x_79; -goto block_53; -} -} -} -else -{ -lean_object* x_80; lean_object* x_81; -lean_dec(x_35); -x_80 = lean_ctor_get(x_54, 0); -lean_inc(x_80); -x_81 = lean_ctor_get(x_54, 1); -lean_inc(x_81); -lean_dec(x_54); -x_38 = x_80; -x_39 = x_81; -goto block_53; -} -block_53: -{ -uint8_t x_40; -x_40 = l_Lean_Exception_isRuntime(x_38); -if (x_40 == 0) -{ -lean_object* x_41; uint8_t x_42; -lean_dec(x_37); -x_41 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_39); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_42 = !lean_is_exclusive(x_41); -if (x_42 == 0) -{ -lean_object* x_43; -x_43 = lean_ctor_get(x_41, 0); -lean_dec(x_43); -lean_ctor_set_tag(x_41, 1); -lean_ctor_set(x_41, 0, x_38); -return x_41; -} -else -{ -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_41, 1); -lean_inc(x_44); -lean_dec(x_41); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_38); -lean_ctor_set(x_45, 1, x_44); -return x_45; -} -} -else -{ -uint8_t x_46; -x_46 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -if (x_46 == 0) -{ -lean_object* x_47; -lean_dec(x_21); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -if (lean_is_scalar(x_37)) { - x_47 = lean_alloc_ctor(1, 2, 0); -} else { - x_47 = x_37; - lean_ctor_set_tag(x_47, 1); -} -lean_ctor_set(x_47, 0, x_38); -lean_ctor_set(x_47, 1, x_39); -return x_47; -} -else -{ -lean_object* x_48; uint8_t x_49; -lean_dec(x_37); -x_48 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_39); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_49 = !lean_is_exclusive(x_48); -if (x_49 == 0) -{ -lean_object* x_50; -x_50 = lean_ctor_get(x_48, 0); -lean_dec(x_50); -lean_ctor_set_tag(x_48, 1); -lean_ctor_set(x_48, 0, x_38); -return x_48; -} -else -{ -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_48, 1); -lean_inc(x_51); -lean_dec(x_48); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_38); -lean_ctor_set(x_52, 1, x_51); -return x_52; -} -} -} -} -} -else -{ -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_110; -x_82 = lean_ctor_get(x_25, 0); -x_83 = lean_ctor_get(x_25, 1); -x_84 = lean_ctor_get(x_25, 2); -x_85 = lean_ctor_get(x_25, 3); -x_86 = lean_ctor_get(x_25, 4); -x_87 = lean_ctor_get(x_25, 6); -lean_inc(x_87); -lean_inc(x_86); -lean_inc(x_85); -lean_inc(x_84); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_25); -x_88 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__6; -x_89 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_89, 0, x_82); -lean_ctor_set(x_89, 1, x_83); -lean_ctor_set(x_89, 2, x_84); -lean_ctor_set(x_89, 3, x_85); -lean_ctor_set(x_89, 4, x_86); -lean_ctor_set(x_89, 5, x_88); -lean_ctor_set(x_89, 6, x_87); -lean_ctor_set(x_24, 1, x_89); -x_90 = lean_st_ref_set(x_8, x_24, x_26); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); -x_92 = l_Lean_Meta_getResetPostponed(x_7, x_8, x_9, x_10, x_91); -x_93 = lean_ctor_get(x_92, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_92, 1); -lean_inc(x_94); -if (lean_is_exclusive(x_92)) { - lean_ctor_release(x_92, 0); - lean_ctor_release(x_92, 1); - x_95 = x_92; -} else { - lean_dec_ref(x_92); - x_95 = lean_box(0); -} -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_110 = lean_is_expr_def_eq(x_1, x_14, x_7, x_8, x_9, x_10, x_94); -if (lean_obj_tag(x_110) == 0) -{ -lean_object* x_111; uint8_t x_112; -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_unbox(x_111); -lean_dec(x_111); -if (x_112 == 0) -{ -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; -lean_dec(x_95); -lean_dec(x_93); -x_113 = lean_ctor_get(x_110, 1); -lean_inc(x_113); -lean_dec(x_110); -x_114 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_113); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_115 = lean_ctor_get(x_114, 1); -lean_inc(x_115); -lean_dec(x_114); -x_116 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_116; -x_16 = x_115; -goto block_19; -} -else -{ -lean_object* x_117; uint8_t x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_110, 1); -lean_inc(x_117); -lean_dec(x_110); -x_118 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_119 = l_Lean_Meta_processPostponed(x_6, x_118, x_7, x_8, x_9, x_10, x_117); -if (lean_obj_tag(x_119) == 0) -{ -lean_object* x_120; uint8_t x_121; -lean_dec(x_95); -x_120 = lean_ctor_get(x_119, 0); -lean_inc(x_120); -x_121 = lean_unbox(x_120); -lean_dec(x_120); -if (x_121 == 0) -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; -lean_dec(x_93); -x_122 = lean_ctor_get(x_119, 1); -lean_inc(x_122); -lean_dec(x_119); -x_123 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_122); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_124 = lean_ctor_get(x_123, 1); -lean_inc(x_124); -lean_dec(x_123); -x_125 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_125; -x_16 = x_124; -goto block_19; -} -else -{ -lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; -lean_dec(x_21); -x_126 = lean_ctor_get(x_119, 1); -lean_inc(x_126); -lean_dec(x_119); -x_127 = l_Lean_Meta_getPostponed___rarg(x_8, x_9, x_10, x_126); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -lean_dec(x_127); -x_130 = l_Lean_PersistentArray_append___rarg(x_93, x_128); -x_131 = l_Lean_Meta_setPostponed(x_130, x_7, x_8, x_9, x_10, x_129); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_132 = lean_ctor_get(x_131, 1); -lean_inc(x_132); -lean_dec(x_131); -x_133 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__8; -x_15 = x_133; -x_16 = x_132; -goto block_19; -} -} -else -{ -lean_object* x_134; lean_object* x_135; -lean_dec(x_93); -x_134 = lean_ctor_get(x_119, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_119, 1); -lean_inc(x_135); -lean_dec(x_119); -x_96 = x_134; -x_97 = x_135; -goto block_109; -} -} -} -else -{ -lean_object* x_136; lean_object* x_137; -lean_dec(x_93); -x_136 = lean_ctor_get(x_110, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_110, 1); -lean_inc(x_137); -lean_dec(x_110); -x_96 = x_136; -x_97 = x_137; -goto block_109; -} -block_109: -{ -uint8_t x_98; -x_98 = l_Lean_Exception_isRuntime(x_96); -if (x_98 == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; -lean_dec(x_95); -x_99 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_97); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -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); -} -if (lean_is_scalar(x_101)) { - x_102 = lean_alloc_ctor(1, 2, 0); -} else { - x_102 = x_101; - lean_ctor_set_tag(x_102, 1); -} -lean_ctor_set(x_102, 0, x_96); -lean_ctor_set(x_102, 1, x_100); -return x_102; -} -else -{ -uint8_t x_103; -x_103 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -if (x_103 == 0) -{ -lean_object* x_104; -lean_dec(x_21); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -if (lean_is_scalar(x_95)) { - x_104 = lean_alloc_ctor(1, 2, 0); -} else { - x_104 = x_95; - lean_ctor_set_tag(x_104, 1); -} -lean_ctor_set(x_104, 0, x_96); -lean_ctor_set(x_104, 1, x_97); -return x_104; -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_95); -x_105 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_97); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_106 = lean_ctor_get(x_105, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_107 = x_105; -} else { - lean_dec_ref(x_105); - x_107 = lean_box(0); -} -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); -} else { - x_108 = x_107; - lean_ctor_set_tag(x_108, 1); -} -lean_ctor_set(x_108, 0, x_96); -lean_ctor_set(x_108, 1, x_106); -return x_108; -} -} -} -} -} -else -{ -lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_172; -x_138 = lean_ctor_get(x_24, 0); -x_139 = lean_ctor_get(x_24, 2); -x_140 = lean_ctor_get(x_24, 3); -x_141 = lean_ctor_get(x_24, 4); -lean_inc(x_141); -lean_inc(x_140); -lean_inc(x_139); -lean_inc(x_138); -lean_dec(x_24); -x_142 = lean_ctor_get(x_25, 0); -lean_inc(x_142); -x_143 = lean_ctor_get(x_25, 1); -lean_inc(x_143); -x_144 = lean_ctor_get(x_25, 2); -lean_inc(x_144); -x_145 = lean_ctor_get(x_25, 3); -lean_inc(x_145); -x_146 = lean_ctor_get(x_25, 4); -lean_inc(x_146); -x_147 = lean_ctor_get(x_25, 6); -lean_inc(x_147); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - lean_ctor_release(x_25, 2); - lean_ctor_release(x_25, 3); - lean_ctor_release(x_25, 4); - lean_ctor_release(x_25, 5); - lean_ctor_release(x_25, 6); - x_148 = x_25; -} else { - lean_dec_ref(x_25); - x_148 = lean_box(0); -} -x_149 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__6; -if (lean_is_scalar(x_148)) { - x_150 = lean_alloc_ctor(0, 7, 0); -} else { - x_150 = x_148; -} -lean_ctor_set(x_150, 0, x_142); -lean_ctor_set(x_150, 1, x_143); -lean_ctor_set(x_150, 2, x_144); -lean_ctor_set(x_150, 3, x_145); -lean_ctor_set(x_150, 4, x_146); -lean_ctor_set(x_150, 5, x_149); -lean_ctor_set(x_150, 6, x_147); -x_151 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_151, 0, x_138); -lean_ctor_set(x_151, 1, x_150); -lean_ctor_set(x_151, 2, x_139); -lean_ctor_set(x_151, 3, x_140); -lean_ctor_set(x_151, 4, x_141); -x_152 = lean_st_ref_set(x_8, x_151, x_26); -x_153 = lean_ctor_get(x_152, 1); -lean_inc(x_153); -lean_dec(x_152); -x_154 = l_Lean_Meta_getResetPostponed(x_7, x_8, x_9, x_10, x_153); -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); -lean_inc(x_156); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_157 = x_154; -} else { - lean_dec_ref(x_154); - x_157 = lean_box(0); -} -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_172 = lean_is_expr_def_eq(x_1, x_14, x_7, x_8, x_9, x_10, x_156); -if (lean_obj_tag(x_172) == 0) -{ -lean_object* x_173; uint8_t x_174; -x_173 = lean_ctor_get(x_172, 0); -lean_inc(x_173); -x_174 = lean_unbox(x_173); -lean_dec(x_173); -if (x_174 == 0) -{ -lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -lean_dec(x_157); -lean_dec(x_155); -x_175 = lean_ctor_get(x_172, 1); -lean_inc(x_175); -lean_dec(x_172); -x_176 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_175); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_177 = lean_ctor_get(x_176, 1); -lean_inc(x_177); -lean_dec(x_176); -x_178 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_178; -x_16 = x_177; -goto block_19; -} -else -{ -lean_object* x_179; uint8_t x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_172, 1); -lean_inc(x_179); -lean_dec(x_172); -x_180 = 0; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -x_181 = l_Lean_Meta_processPostponed(x_6, x_180, x_7, x_8, x_9, x_10, x_179); -if (lean_obj_tag(x_181) == 0) -{ -lean_object* x_182; uint8_t x_183; -lean_dec(x_157); -x_182 = lean_ctor_get(x_181, 0); -lean_inc(x_182); -x_183 = lean_unbox(x_182); -lean_dec(x_182); -if (x_183 == 0) -{ -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_155); -x_184 = lean_ctor_get(x_181, 1); -lean_inc(x_184); -lean_dec(x_181); -x_185 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_184); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_186 = lean_ctor_get(x_185, 1); -lean_inc(x_186); -lean_dec(x_185); -x_187 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__7; -x_15 = x_187; -x_16 = x_186; -goto block_19; -} -else -{ -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_dec(x_21); -x_188 = lean_ctor_get(x_181, 1); -lean_inc(x_188); -lean_dec(x_181); -x_189 = l_Lean_Meta_getPostponed___rarg(x_8, x_9, x_10, x_188); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -lean_dec(x_189); -x_192 = l_Lean_PersistentArray_append___rarg(x_155, x_190); -x_193 = l_Lean_Meta_setPostponed(x_192, x_7, x_8, x_9, x_10, x_191); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_194 = lean_ctor_get(x_193, 1); -lean_inc(x_194); -lean_dec(x_193); -x_195 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__2___closed__8; -x_15 = x_195; -x_16 = x_194; -goto block_19; -} -} -else -{ -lean_object* x_196; lean_object* x_197; -lean_dec(x_155); -x_196 = lean_ctor_get(x_181, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_181, 1); -lean_inc(x_197); -lean_dec(x_181); -x_158 = x_196; -x_159 = x_197; -goto block_171; -} -} -} -else -{ -lean_object* x_198; lean_object* x_199; -lean_dec(x_155); -x_198 = lean_ctor_get(x_172, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_172, 1); -lean_inc(x_199); -lean_dec(x_172); -x_158 = x_198; -x_159 = x_199; -goto block_171; -} -block_171: -{ -uint8_t x_160; -x_160 = l_Lean_Exception_isRuntime(x_158); -if (x_160 == 0) -{ -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; -lean_dec(x_157); -x_161 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_159); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_162 = lean_ctor_get(x_161, 1); -lean_inc(x_162); -if (lean_is_exclusive(x_161)) { - lean_ctor_release(x_161, 0); - lean_ctor_release(x_161, 1); - x_163 = x_161; -} else { - lean_dec_ref(x_161); - x_163 = lean_box(0); -} -if (lean_is_scalar(x_163)) { - x_164 = lean_alloc_ctor(1, 2, 0); -} else { - x_164 = x_163; - lean_ctor_set_tag(x_164, 1); -} -lean_ctor_set(x_164, 0, x_158); -lean_ctor_set(x_164, 1, x_162); -return x_164; -} -else -{ -uint8_t x_165; -x_165 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -if (x_165 == 0) -{ -lean_object* x_166; -lean_dec(x_21); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -if (lean_is_scalar(x_157)) { - x_166 = lean_alloc_ctor(1, 2, 0); -} else { - x_166 = x_157; - lean_ctor_set_tag(x_166, 1); -} -lean_ctor_set(x_166, 0, x_158); -lean_ctor_set(x_166, 1, x_159); -return x_166; -} -else -{ -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_157); -x_167 = l_Lean_Meta_SavedState_restore(x_21, x_7, x_8, x_9, x_10, x_159); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_21); -x_168 = lean_ctor_get(x_167, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_167)) { - lean_ctor_release(x_167, 0); - lean_ctor_release(x_167, 1); - x_169 = x_167; -} else { - lean_dec_ref(x_167); - x_169 = lean_box(0); -} -if (lean_is_scalar(x_169)) { - x_170 = lean_alloc_ctor(1, 2, 0); -} else { - x_170 = x_169; - lean_ctor_set_tag(x_170, 1); -} -lean_ctor_set(x_170, 0, x_158); -lean_ctor_set(x_170, 1, x_168); -return x_170; -} -} -} -} -block_19: -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -return x_18; -} -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta(lean_object* x_1, 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: { @@ -5375,7 +4534,7 @@ x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); if (lean_obj_tag(x_17) == 7) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; +lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); @@ -5385,67 +4544,69 @@ x_20 = lean_ctor_get(x_17, 1); lean_inc(x_20); x_21 = lean_ctor_get_uint8(x_17, sizeof(void*)*3 + 8); lean_dec(x_17); -x_22 = 1; -x_23 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1(x_1, x_2, x_19, x_20, x_21, x_22, x_3, x_4, x_5, x_6, x_18); -if (lean_obj_tag(x_23) == 0) +x_22 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1; +x_23 = l_Lean_Expr_app___override(x_2, x_22); +x_24 = l_Lean_Expr_lam___override(x_19, x_20, x_23, x_21); +x_25 = lean_is_expr_def_eq(x_1, x_24, x_3, x_4, x_5, x_6, x_18); +if (lean_obj_tag(x_25) == 0) { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +uint8_t x_26; +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) { -lean_object* x_25; uint8_t x_26; uint8_t x_27; lean_object* x_28; -x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_unbox(x_25); -lean_dec(x_25); -x_27 = l_Bool_toLBool(x_26); -x_28 = lean_box(x_27); -lean_ctor_set(x_23, 0, x_28); -return x_23; +lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_unbox(x_27); +lean_dec(x_27); +x_29 = l_Bool_toLBool(x_28); +x_30 = lean_box(x_29); +lean_ctor_set(x_25, 0, x_30); +return x_25; } else { -lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; lean_object* x_34; -x_29 = lean_ctor_get(x_23, 0); -x_30 = lean_ctor_get(x_23, 1); -lean_inc(x_30); -lean_inc(x_29); -lean_dec(x_23); -x_31 = lean_unbox(x_29); -lean_dec(x_29); -x_32 = l_Bool_toLBool(x_31); -x_33 = lean_box(x_32); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_30); -return x_34; +lean_object* x_31; lean_object* x_32; uint8_t x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; +x_31 = lean_ctor_get(x_25, 0); +x_32 = lean_ctor_get(x_25, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_25); +x_33 = lean_unbox(x_31); +lean_dec(x_31); +x_34 = l_Bool_toLBool(x_33); +x_35 = lean_box(x_34); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_32); +return x_36; } } else { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_23); -if (x_35 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_25); +if (x_37 == 0) { -return x_23; +return x_25; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_23, 0); -x_37 = lean_ctor_get(x_23, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_23); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_25, 0); +x_39 = lean_ctor_get(x_25, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_25); +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 { -uint8_t x_39; +uint8_t x_41; lean_dec(x_17); lean_dec(x_6); lean_dec(x_5); @@ -5453,119 +4614,107 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_39 = !lean_is_exclusive(x_16); -if (x_39 == 0) +x_41 = !lean_is_exclusive(x_16); +if (x_41 == 0) { -lean_object* x_40; uint8_t x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_16, 0); -lean_dec(x_40); -x_41 = 2; -x_42 = lean_box(x_41); -lean_ctor_set(x_16, 0, x_42); +lean_object* x_42; uint8_t x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_16, 0); +lean_dec(x_42); +x_43 = 2; +x_44 = lean_box(x_43); +lean_ctor_set(x_16, 0, x_44); return x_16; } else { -lean_object* x_43; uint8_t x_44; lean_object* x_45; lean_object* x_46; -x_43 = lean_ctor_get(x_16, 1); -lean_inc(x_43); +lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_16, 1); +lean_inc(x_45); lean_dec(x_16); -x_44 = 2; -x_45 = lean_box(x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -return x_46; +x_46 = 2; +x_47 = lean_box(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_45); +return x_48; } } } else { -uint8_t x_47; +uint8_t x_49; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_47 = !lean_is_exclusive(x_16); -if (x_47 == 0) +x_49 = !lean_is_exclusive(x_16); +if (x_49 == 0) { return x_16; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_16, 0); -x_49 = lean_ctor_get(x_16, 1); -lean_inc(x_49); -lean_inc(x_48); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_16, 0); +x_51 = lean_ctor_get(x_16, 1); +lean_inc(x_51); +lean_inc(x_50); lean_dec(x_16); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_51; +uint8_t x_53; 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); -if (x_51 == 0) +x_53 = !lean_is_exclusive(x_13); +if (x_53 == 0) { return x_13; } 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); -lean_inc(x_53); -lean_inc(x_52); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_13, 0); +x_55 = lean_ctor_get(x_13, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_13); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } else { -uint8_t x_55; lean_object* x_56; lean_object* x_57; +uint8_t x_57; lean_object* x_58; lean_object* x_59; 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 = 2; -x_56 = lean_box(x_55); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_56); -lean_ctor_set(x_57, 1, x_7); -return x_57; -} -} +x_57 = 2; +x_58 = lean_box(x_57); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_7); +return x_59; } } -LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___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; uint8_t x_13; lean_object* x_14; -x_12 = lean_unbox(x_5); -lean_dec(x_5); -x_13 = lean_unbox(x_6); -lean_dec(x_6); -x_14 = l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1(x_1, x_2, x_3, x_4, x_12, x_13, x_7, x_8, x_9, x_10, x_11); -return x_14; } } LEAN_EXPORT lean_object* l_Lean_Meta_isDefEqNative(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { @@ -11924,87 +11073,21 @@ lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_inc(x_9); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_4, x_7, x_8, x_9, x_10, x_13); lean_dec(x_9); return x_14; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -12012,16 +11095,16 @@ x_1 = lean_mk_string_from_bytes(" ", 1); return x_1; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1; +x_1 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static double _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3() { +static double _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -12031,7 +11114,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4() { _start: { lean_object* x_1; @@ -12039,7 +11122,7 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, double x_10, double 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_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, uint8_t x_9, double x_10, double 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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; double x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; @@ -12050,7 +11133,7 @@ x_20 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_21 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_21, 0, x_20); lean_ctor_set(x_21, 1, x_19); -x_22 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_22 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -12060,7 +11143,7 @@ lean_ctor_set(x_24, 1, x_2); 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_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_26 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3; lean_inc(x_5); lean_inc(x_3); x_27 = lean_alloc_ctor(0, 2, 17); @@ -12069,7 +11152,7 @@ lean_ctor_set(x_27, 1, x_5); lean_ctor_set_float(x_27, sizeof(void*)*2, x_26); lean_ctor_set_float(x_27, sizeof(void*)*2 + 8, x_26); lean_ctor_set_uint8(x_27, sizeof(void*)*2 + 16, x_4); -x_28 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; +x_28 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4; x_29 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_8, x_28); if (x_29 == 0) { @@ -12079,7 +11162,7 @@ lean_object* x_30; lean_object* x_31; lean_dec(x_5); lean_dec(x_3); x_30 = lean_box(0); -x_31 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_6, x_7, x_25, x_1, x_27, x_30, x_13, x_14, x_15, x_16, x_17); +x_31 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(x_6, x_7, x_25, x_1, x_27, x_30, x_13, x_14, x_15, x_16, x_17); lean_dec(x_1); return x_31; } @@ -12094,7 +11177,7 @@ lean_ctor_set_float(x_32, sizeof(void*)*2, x_10); lean_ctor_set_float(x_32, sizeof(void*)*2 + 8, x_11); lean_ctor_set_uint8(x_32, sizeof(void*)*2 + 16, x_4); x_33 = lean_box(0); -x_34 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_6, x_7, x_25, x_1, x_32, x_33, x_13, x_14, x_15, x_16, x_17); +x_34 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(x_6, x_7, x_25, x_1, x_32, x_33, x_13, x_14, x_15, x_16, x_17); lean_dec(x_1); return x_34; } @@ -12110,13 +11193,21 @@ lean_ctor_set_float(x_35, sizeof(void*)*2, x_10); lean_ctor_set_float(x_35, sizeof(void*)*2 + 8, x_11); lean_ctor_set_uint8(x_35, sizeof(void*)*2 + 16, x_4); x_36 = lean_box(0); -x_37 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_6, x_7, x_25, x_1, x_35, x_36, x_13, x_14, x_15, x_16, x_17); +x_37 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(x_6, x_7, x_25, x_1, x_35, x_36, x_13, x_14, x_15, x_16, x_17); lean_dec(x_1); return x_37; } } } -static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_useHearbeats; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -12124,7 +11215,37 @@ x_1 = l_Lean_trace_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, 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) { +static lean_object* _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_threshold; +return x_1; +} +} +static double _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +static double _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, uint8_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; @@ -12181,7 +11302,7 @@ lean_inc(x_9); x_32 = lean_apply_5(x_1, x_9, x_10, x_31, x_12, x_16); if (lean_obj_tag(x_32) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint8_t x_39; x_33 = lean_ctor_get(x_32, 0); lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); @@ -12194,33 +11315,111 @@ lean_inc(x_36); x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); lean_dec(x_35); -x_38 = lean_alloc_closure((void*)(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1), 6, 1); -lean_closure_set(x_38, 0, x_2); +x_38 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1; +x_39 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_38); +if (x_39 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_io_mono_nanos_now(x_37); +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_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_39 = l_Lean_withStartStopSeconds___at_Lean_Meta_processPostponed___spec__3(x_38, x_9, x_10, x_11, x_12, x_37); -if (lean_obj_tag(x_39) == 0) +x_111 = lean_apply_5(x_7, x_9, x_10, x_11, x_12, x_110); +if (lean_obj_tag(x_111) == 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; uint8_t x_46; lean_object* x_93; uint8_t x_94; -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -x_42 = lean_ctor_get(x_39, 1); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; uint8_t x_118; lean_object* x_119; double x_120; double x_121; double x_122; double x_123; double x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_114, 0, x_112); +x_115 = lean_io_mono_nanos_now(x_113); +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = 0; +x_119 = lean_unsigned_to_nat(0u); +x_120 = l_Float_ofScientific(x_109, x_118, x_119); +lean_dec(x_109); +x_121 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; +x_122 = lean_float_div(x_120, x_121); +x_123 = l_Float_ofScientific(x_116, x_118, x_119); +lean_dec(x_116); +x_124 = lean_float_div(x_123, x_121); +x_125 = lean_box_float(x_122); +x_126 = lean_box_float(x_124); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_114); +lean_ctor_set(x_128, 1, x_127); +x_40 = x_128; +x_41 = x_117; +goto block_107; +} +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; uint8_t x_135; lean_object* x_136; double x_137; double x_138; double x_139; double x_140; double x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +x_129 = lean_ctor_get(x_111, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_111, 1); +lean_inc(x_130); +lean_dec(x_111); +x_131 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_131, 0, x_129); +x_132 = lean_io_mono_nanos_now(x_130); +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +x_134 = lean_ctor_get(x_132, 1); +lean_inc(x_134); +lean_dec(x_132); +x_135 = 0; +x_136 = lean_unsigned_to_nat(0u); +x_137 = l_Float_ofScientific(x_109, x_135, x_136); +lean_dec(x_109); +x_138 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5; +x_139 = lean_float_div(x_137, x_138); +x_140 = l_Float_ofScientific(x_133, x_135, x_136); +lean_dec(x_133); +x_141 = lean_float_div(x_140, x_138); +x_142 = lean_box_float(x_139); +x_143 = lean_box_float(x_141); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_142); +lean_ctor_set(x_144, 1, x_143); +x_145 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_145, 0, x_131); +lean_ctor_set(x_145, 1, x_144); +x_40 = x_145; +x_41 = x_134; +goto block_107; +} +block_107: +{ +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_93; uint8_t x_94; +x_42 = lean_ctor_get(x_40, 1); lean_inc(x_42); -lean_dec(x_39); x_43 = lean_ctor_get(x_40, 0); lean_inc(x_43); lean_dec(x_40); -x_44 = lean_ctor_get(x_41, 0); +x_44 = lean_ctor_get(x_42, 0); lean_inc(x_44); -x_45 = lean_ctor_get(x_41, 1); +x_45 = lean_ctor_get(x_42, 1); lean_inc(x_45); -lean_dec(x_41); -x_93 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; -x_94 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_6, x_93); +lean_dec(x_42); +x_93 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_94 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_93); if (x_94 == 0) { uint8_t x_95; @@ -12230,18 +11429,25 @@ goto block_92; } else { -double x_96; double x_97; double x_98; double x_99; uint8_t x_100; -x_96 = l_Lean_trace_profiler_threshold_getSecs(x_6); -x_97 = lean_unbox_float(x_45); -x_98 = lean_unbox_float(x_44); -x_99 = lean_float_sub(x_97, x_98); -x_100 = lean_float_decLt(x_96, x_99); -x_46 = x_100; +double x_96; double x_97; double x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; double x_103; double x_104; double x_105; uint8_t x_106; +x_96 = lean_unbox_float(x_45); +x_97 = lean_unbox_float(x_44); +x_98 = lean_float_sub(x_96, x_97); +x_99 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_100 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_99); +x_101 = 0; +x_102 = lean_unsigned_to_nat(0u); +x_103 = l_Float_ofScientific(x_100, x_101, x_102); +lean_dec(x_100); +x_104 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4; +x_105 = lean_float_div(x_103, x_104); +x_106 = lean_float_decLt(x_105, x_98); +x_46 = x_106; goto block_92; } block_92: { -if (x_7 == 0) +if (x_6 == 0) { if (x_46 == 0) { @@ -12250,10 +11456,10 @@ lean_dec(x_45); lean_dec(x_44); lean_dec(x_36); lean_dec(x_22); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_47 = lean_st_ref_take(x_12, x_42); +lean_dec(x_2); +x_47 = lean_st_ref_take(x_12, x_41); x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); x_49 = lean_ctor_get(x_47, 1); @@ -12270,7 +11476,7 @@ x_53 = lean_st_ref_set(x_12, x_48, x_49); x_54 = lean_ctor_get(x_53, 1); lean_inc(x_54); lean_dec(x_53); -x_55 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_43, x_9, x_10, x_11, x_12, x_54); +x_55 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_43, x_9, x_10, x_11, x_12, x_54); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -12352,7 +11558,7 @@ x_73 = lean_st_ref_set(x_12, x_72, x_49); x_74 = lean_ctor_get(x_73, 1); lean_inc(x_74); lean_dec(x_73); -x_75 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__5(x_43, x_9, x_10, x_11, x_12, x_74); +x_75 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_43, x_9, x_10, x_11, x_12, x_74); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -12416,11 +11622,11 @@ x_85 = lean_unbox_float(x_44); lean_dec(x_44); x_86 = lean_unbox_float(x_45); lean_dec(x_45); -x_87 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_43, x_36, x_3, x_4, x_5, x_15, x_22, x_6, x_46, x_85, x_86, x_84, x_9, x_10, x_11, x_12, x_42); +x_87 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_43, x_36, x_3, x_4, x_5, x_15, x_22, x_2, x_46, x_85, x_86, x_84, x_9, x_10, x_11, x_12, x_41); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_2); return x_87; } } @@ -12432,51 +11638,343 @@ x_89 = lean_unbox_float(x_44); lean_dec(x_44); x_90 = lean_unbox_float(x_45); lean_dec(x_45); -x_91 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_43, x_36, x_3, x_4, x_5, x_15, x_22, x_6, x_46, x_89, x_90, x_88, x_9, x_10, x_11, x_12, x_42); +x_91 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_43, x_36, x_3, x_4, x_5, x_15, x_22, x_2, x_46, x_89, x_90, x_88, x_9, x_10, x_11, x_12, x_41); lean_dec(x_12); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_2); return x_91; } } } +} else { -uint8_t x_101; +lean_object* x_146; lean_object* x_147; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_212 = lean_io_get_num_heartbeats(x_37); +x_213 = lean_ctor_get(x_212, 0); +lean_inc(x_213); +x_214 = lean_ctor_get(x_212, 1); +lean_inc(x_214); +lean_dec(x_212); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_215 = lean_apply_5(x_7, x_9, x_10, x_11, x_12, x_214); +if (lean_obj_tag(x_215) == 0) +{ +lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; uint8_t x_222; lean_object* x_223; double x_224; double x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_218, 0, x_216); +x_219 = lean_io_get_num_heartbeats(x_217); +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +x_221 = lean_ctor_get(x_219, 1); +lean_inc(x_221); +lean_dec(x_219); +x_222 = 0; +x_223 = lean_unsigned_to_nat(0u); +x_224 = l_Float_ofScientific(x_213, x_222, x_223); +lean_dec(x_213); +x_225 = l_Float_ofScientific(x_220, x_222, x_223); +lean_dec(x_220); +x_226 = lean_box_float(x_224); +x_227 = lean_box_float(x_225); +x_228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_228, 0, x_226); +lean_ctor_set(x_228, 1, x_227); +x_229 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_229, 0, x_218); +lean_ctor_set(x_229, 1, x_228); +x_146 = x_229; +x_147 = x_221; +goto block_211; +} +else +{ +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; uint8_t x_236; lean_object* x_237; double x_238; double x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; +x_230 = lean_ctor_get(x_215, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_215, 1); +lean_inc(x_231); +lean_dec(x_215); +x_232 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_232, 0, x_230); +x_233 = lean_io_get_num_heartbeats(x_231); +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +x_236 = 0; +x_237 = lean_unsigned_to_nat(0u); +x_238 = l_Float_ofScientific(x_213, x_236, x_237); +lean_dec(x_213); +x_239 = l_Float_ofScientific(x_234, x_236, x_237); +lean_dec(x_234); +x_240 = lean_box_float(x_238); +x_241 = lean_box_float(x_239); +x_242 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_242, 0, x_240); +lean_ctor_set(x_242, 1, x_241); +x_243 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_243, 0, x_232); +lean_ctor_set(x_243, 1, x_242); +x_146 = x_243; +x_147 = x_235; +goto block_211; +} +block_211: +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; uint8_t x_152; lean_object* x_199; uint8_t x_200; +x_148 = lean_ctor_get(x_146, 1); +lean_inc(x_148); +x_149 = lean_ctor_get(x_146, 0); +lean_inc(x_149); +lean_dec(x_146); +x_150 = lean_ctor_get(x_148, 0); +lean_inc(x_150); +x_151 = lean_ctor_get(x_148, 1); +lean_inc(x_151); +lean_dec(x_148); +x_199 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_200 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_2, x_199); +if (x_200 == 0) +{ +uint8_t x_201; +x_201 = 0; +x_152 = x_201; +goto block_198; +} +else +{ +double x_202; double x_203; double x_204; lean_object* x_205; lean_object* x_206; uint8_t x_207; lean_object* x_208; double x_209; uint8_t x_210; +x_202 = lean_unbox_float(x_151); +x_203 = lean_unbox_float(x_150); +x_204 = lean_float_sub(x_202, x_203); +x_205 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_206 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_2, x_205); +x_207 = 0; +x_208 = lean_unsigned_to_nat(0u); +x_209 = l_Float_ofScientific(x_206, x_207, x_208); +lean_dec(x_206); +x_210 = lean_float_decLt(x_209, x_204); +x_152 = x_210; +goto block_198; +} +block_198: +{ +if (x_6 == 0) +{ +if (x_152 == 0) +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; uint8_t x_156; +lean_dec(x_151); +lean_dec(x_150); lean_dec(x_36); lean_dec(x_22); -lean_dec(x_15); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +x_153 = lean_st_ref_take(x_12, x_147); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +x_156 = !lean_is_exclusive(x_154); +if (x_156 == 0) +{ +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_157 = lean_ctor_get(x_154, 3); +x_158 = l_Lean_PersistentArray_append___rarg(x_15, x_157); +lean_ctor_set(x_154, 3, x_158); +x_159 = lean_st_ref_set(x_12, x_154, x_155); +x_160 = lean_ctor_get(x_159, 1); +lean_inc(x_160); +lean_dec(x_159); +x_161 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_149, x_9, x_10, x_11, x_12, x_160); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_101 = !lean_is_exclusive(x_39); -if (x_101 == 0) +lean_dec(x_149); +if (lean_obj_tag(x_161) == 0) { -return x_39; +uint8_t x_162; +x_162 = !lean_is_exclusive(x_161); +if (x_162 == 0) +{ +return x_161; } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_39, 0); -x_103 = lean_ctor_get(x_39, 1); -lean_inc(x_103); -lean_inc(x_102); -lean_dec(x_39); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -return x_104; +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_161, 0); +x_164 = lean_ctor_get(x_161, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_161); +x_165 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_165, 0, x_163); +lean_ctor_set(x_165, 1, x_164); +return x_165; } } +else +{ +uint8_t x_166; +x_166 = !lean_is_exclusive(x_161); +if (x_166 == 0) +{ +return x_161; } else { -uint8_t x_105; +lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_167 = lean_ctor_get(x_161, 0); +x_168 = lean_ctor_get(x_161, 1); +lean_inc(x_168); +lean_inc(x_167); +lean_dec(x_161); +x_169 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_169, 0, x_167); +lean_ctor_set(x_169, 1, x_168); +return x_169; +} +} +} +else +{ +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_170 = lean_ctor_get(x_154, 0); +x_171 = lean_ctor_get(x_154, 1); +x_172 = lean_ctor_get(x_154, 2); +x_173 = lean_ctor_get(x_154, 3); +x_174 = lean_ctor_get(x_154, 4); +x_175 = lean_ctor_get(x_154, 5); +x_176 = lean_ctor_get(x_154, 6); +lean_inc(x_176); +lean_inc(x_175); +lean_inc(x_174); +lean_inc(x_173); +lean_inc(x_172); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_154); +x_177 = l_Lean_PersistentArray_append___rarg(x_15, x_173); +x_178 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_178, 0, x_170); +lean_ctor_set(x_178, 1, x_171); +lean_ctor_set(x_178, 2, x_172); +lean_ctor_set(x_178, 3, x_177); +lean_ctor_set(x_178, 4, x_174); +lean_ctor_set(x_178, 5, x_175); +lean_ctor_set(x_178, 6, x_176); +x_179 = lean_st_ref_set(x_12, x_178, x_155); +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +lean_dec(x_179); +x_181 = l_MonadExcept_ofExcept___at_Lean_Meta_processPostponed___spec__4(x_149, x_9, x_10, x_11, x_12, x_180); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_149); +if (lean_obj_tag(x_181) == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_184 = x_181; +} else { + lean_dec_ref(x_181); + x_184 = lean_box(0); +} +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(0, 2, 0); +} else { + x_185 = x_184; +} +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +return x_185; +} +else +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_186 = lean_ctor_get(x_181, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_181, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_188 = x_181; +} else { + lean_dec_ref(x_181); + x_188 = lean_box(0); +} +if (lean_is_scalar(x_188)) { + x_189 = lean_alloc_ctor(1, 2, 0); +} else { + x_189 = x_188; +} +lean_ctor_set(x_189, 0, x_186); +lean_ctor_set(x_189, 1, x_187); +return x_189; +} +} +} +else +{ +lean_object* x_190; double x_191; double x_192; lean_object* x_193; +x_190 = lean_box(0); +x_191 = lean_unbox_float(x_150); +lean_dec(x_150); +x_192 = lean_unbox_float(x_151); +lean_dec(x_151); +x_193 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_149, x_36, x_3, x_4, x_5, x_15, x_22, x_2, x_152, x_191, x_192, x_190, x_9, x_10, x_11, x_12, x_147); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +return x_193; +} +} +else +{ +lean_object* x_194; double x_195; double x_196; lean_object* x_197; +x_194 = lean_box(0); +x_195 = lean_unbox_float(x_150); +lean_dec(x_150); +x_196 = lean_unbox_float(x_151); +lean_dec(x_151); +x_197 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_149, x_36, x_3, x_4, x_5, x_15, x_22, x_2, x_152, x_195, x_196, x_194, x_9, x_10, x_11, x_12, x_147); +lean_dec(x_12); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_2); +return x_197; +} +} +} +} +} +else +{ +uint8_t x_244; lean_dec(x_31); lean_dec(x_22); lean_dec(x_15); @@ -12484,27 +11982,27 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); -lean_dec(x_6); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); -x_105 = !lean_is_exclusive(x_32); -if (x_105 == 0) +x_244 = !lean_is_exclusive(x_32); +if (x_244 == 0) { return x_32; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_32, 0); -x_107 = lean_ctor_get(x_32, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_245; lean_object* x_246; lean_object* x_247; +x_245 = lean_ctor_get(x_32, 0); +x_246 = lean_ctor_get(x_32, 1); +lean_inc(x_246); +lean_inc(x_245); lean_dec(x_32); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +x_247 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_247, 0, x_245); +lean_ctor_set(x_247, 1, x_246); +return x_247; } } } @@ -12526,7 +12024,7 @@ lean_object* x_15; lean_object* x_16; uint8_t x_17; x_15 = lean_ctor_get(x_12, 1); lean_inc(x_15); lean_dec(x_12); -x_16 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1; +x_16 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; x_17 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_16); if (x_17 == 0) { @@ -12588,7 +12086,7 @@ lean_object* x_27; uint8_t x_28; lean_object* x_29; x_27 = lean_box(0); x_28 = lean_unbox(x_13); lean_dec(x_13); -x_29 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_2, x_3, x_1, x_4, x_5, x_11, x_28, x_27, x_6, x_7, x_8, x_9, x_15); +x_29 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_2, x_11, x_1, x_4, x_5, x_28, x_3, x_27, x_6, x_7, x_8, x_9, x_15); return x_29; } } @@ -12601,7 +12099,7 @@ lean_dec(x_12); x_31 = lean_box(0); x_32 = lean_unbox(x_13); lean_dec(x_13); -x_33 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_2, x_3, x_1, x_4, x_5, x_11, x_32, x_31, x_6, x_7, x_8, x_9, x_30); +x_33 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_2, x_11, x_1, x_4, x_5, x_32, x_3, x_31, x_6, x_7, x_8, x_9, x_30); return x_33; } } @@ -13377,11 +12875,11 @@ return x_19; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___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) { _start: { lean_object* x_12; -x_12 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_12 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -13390,7 +12888,7 @@ lean_dec(x_4); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -13419,7 +12917,7 @@ x_20 = lean_unbox_float(x_10); lean_dec(x_10); x_21 = lean_unbox_float(x_11); lean_dec(x_11); -x_22 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_19, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17); +x_22 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_19, x_20, x_21, x_12, x_13, x_14, x_15, x_16, x_17); lean_dec(x_16); lean_dec(x_14); lean_dec(x_13); @@ -13428,15 +12926,15 @@ lean_dec(x_8); return x_22; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { _start: { uint8_t x_14; uint8_t x_15; lean_object* x_16; x_14 = lean_unbox(x_4); lean_dec(x_4); -x_15 = lean_unbox(x_7); -lean_dec(x_7); -x_16 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4(x_1, x_2, x_3, x_14, x_5, x_6, x_15, x_8, x_9, x_10, x_11, x_12, x_13); +x_15 = lean_unbox(x_6); +lean_dec(x_6); +x_16 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3(x_1, x_2, x_3, x_14, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_16; } } @@ -17364,7 +16862,7 @@ lean_dec(x_5); return x_10; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1() { _start: { lean_object* x_1; @@ -17372,26 +16870,26 @@ x_1 = lean_mk_string_from_bytes("checkAssignment", 15); return x_1; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2; x_3 = l_Lean_registerInternalExceptionId(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1() { _start: { lean_object* x_1; @@ -17399,21 +16897,21 @@ x_1 = lean_mk_string_from_bytes("outOfScope", 10); return x_1; } } -static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2() { +static lean_object* _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2; +x_2 = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2; x_3 = l_Lean_registerInternalExceptionId(x_2, x_1); return x_3; } @@ -17724,7 +17222,7 @@ x_15 = l_Lean_MessageData_ofExpr(x_14); x_16 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_16, 0, x_12); lean_ctor_set(x_16, 1, x_15); -x_17 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_17 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_18 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_18, 0, x_16); lean_ctor_set(x_18, 1, x_17); @@ -18227,7 +17725,7 @@ if (x_17 == 0) { lean_object* x_18; double 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; uint8_t x_28; x_18 = lean_ctor_get(x_15, 3); -x_19 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_19 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3; x_20 = 0; x_21 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__8; x_22 = lean_alloc_ctor(0, 2, 17); @@ -18289,7 +17787,7 @@ lean_inc(x_36); lean_inc(x_35); lean_inc(x_34); lean_dec(x_15); -x_41 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3; +x_41 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3; x_42 = 0; x_43 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___spec__1___lambda__2___closed__8; x_44 = lean_alloc_ctor(0, 2, 17); @@ -54471,7 +53969,7 @@ x_30 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_31 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); -x_32 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_32 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_33 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_33, 0, x_31); lean_ctor_set(x_33, 1, x_32); @@ -55542,7 +55040,7 @@ x_22 = l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_processConstApprox_go___clos x_23 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_23, 0, x_22); lean_ctor_set(x_23, 1, x_21); -x_24 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_24 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_25 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_25, 0, x_23); lean_ctor_set(x_25, 1, x_24); @@ -56198,7 +55696,7 @@ x_19 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_19); lean_ctor_set(x_20, 1, x_18); -x_21 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_21 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___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); @@ -56977,7 +56475,7 @@ x_49 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -67955,7 +67453,7 @@ x_49 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_is x_50 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_50, 0, x_49); lean_ctor_set(x_50, 1, x_48); -x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2; +x_51 = l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2; x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); @@ -82150,7 +81648,7 @@ lean_dec(x_4); return x_11; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -82160,17 +81658,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_ExprDefEq___hyg_19254____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____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_ExprDefEq___hyg_19254____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__3() { _start: { lean_object* x_1; @@ -82178,17 +81676,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____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_ExprDefEq___hyg_19254____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__5() { _start: { lean_object* x_1; @@ -82196,37 +81694,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____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_ExprDefEq___hyg_19254____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___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_ExprDefEq___hyg_19254____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7; x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_6____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_ExprDefEq___hyg_19254____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__9() { _start: { lean_object* x_1; @@ -82234,17 +81732,17 @@ x_1 = lean_mk_string_from_bytes("ExprDefEq", 9); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____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_ExprDefEq___hyg_19254____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__11() { _start: { lean_object* x_1; @@ -82252,33 +81750,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____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_ExprDefEq___hyg_19254____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12; -x_2 = lean_unsigned_to_nat(19254u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12; +x_2 = lean_unsigned_to_nat(19212u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Std_Range_forIn_loop___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___spec__1___closed__1; x_3 = 0; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__13; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -82973,8 +82471,8 @@ l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__6 = _ lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__6); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__7 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__7(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEtaStruct_go___closed__7); -l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1 = _init_l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_checkpointDefEq___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___spec__1___closed__1); +l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1(); +lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqEta___closed__1); l_Lean_Meta_isDefEqStringLit___closed__1 = _init_l_Lean_Meta_isDefEqStringLit___closed__1(); lean_mark_persistent(l_Lean_Meta_isDefEqStringLit___closed__1); l_Lean_Meta_isDefEqStringLit___closed__2 = _init_l_Lean_Meta_isDefEqStringLit___closed__2(); @@ -82995,15 +82493,21 @@ l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2 = lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__2); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_isDefEqArgsFirstPass___closed__3); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__1); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__2); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__3(); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__2___closed__4); l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1(); lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__1); l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2(); lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__2); l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__3); l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__4); -l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__4___closed__1); +l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5 = _init_l_Lean_withTraceNodeBefore___at___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___spec__1___lambda__3___closed__5(); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__2 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_checkTypesAndAssign___lambda__1___closed__2(); @@ -83038,20 +82542,20 @@ l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDecl lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__1); l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2 = _init_l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2(); lean_mark_persistent(l___private_Lean_Meta_ExprDefEq_0__Lean_Meta_mkLambdaFVarsWithLetDeps_hasLetDeclsInBetween___closed__2); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__1); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784____closed__2); -if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5784_(lean_io_mk_world()); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__1); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781____closed__2); +if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5781_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_CheckAssignment_checkAssignmentExceptionId = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_CheckAssignment_checkAssignmentExceptionId); lean_dec_ref(res); -}l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__1); -l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2(); -lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803____closed__2); -if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5803_(lean_io_mk_world()); +}l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__1); +l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2 = _init_l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2(); +lean_mark_persistent(l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800____closed__2); +if (builtin) {res = l_Lean_Meta_CheckAssignment_initFn____x40_Lean_Meta_ExprDefEq___hyg_5800_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_CheckAssignment_outOfScopeExceptionId = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_CheckAssignment_outOfScopeExceptionId); @@ -83274,33 +82778,33 @@ l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__19 = _init_l_Lean_Meta_isEx lean_mark_persistent(l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__19); l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20 = _init_l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20(); lean_mark_persistent(l_Lean_Meta_isExprDefEqAuxImpl___lambda__3___closed__20); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254____closed__13); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19254_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212____closed__13); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_ExprDefEq___hyg_19212_(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/IndPredBelow.c b/stage0/stdlib/Lean/Meta/IndPredBelow.c index af7336eda461..183c30e6947a 100644 --- a/stage0/stdlib/Lean/Meta/IndPredBelow.c +++ b/stage0/stdlib/Lean/Meta/IndPredBelow.c @@ -60,7 +60,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at_Lean_Meta_ static lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_mkBelowMatcher___spec__3___lambda__2___closed__5; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_IndPredBelow_backwardsChaining___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__4(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_proveBrecOn___closed__3; @@ -136,10 +135,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_addBelowPattern LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__7___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_checkCount___closed__3; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at___private_Lean_Meta_IndPredBelow_0__Lean_Meta_IndPredBelow_belowType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_IndPredBelow_mkBelow___spec__2___closed__2; static lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_mkBelowMatcher___spec__3___lambda__2___closed__6; -lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_RecDepth___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkCasesOn___at_Lean_Meta_IndPredBelow_mkBelow___spec__1___closed__8; lean_object* l_Lean_stringToMessageData(lean_object*); @@ -178,6 +175,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_proveBrecOn_applyIH(lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkConstructor(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*); lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_IndPredBelow_backwardsChaining___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_mkCasesOn___at_Lean_Meta_IndPredBelow_mkBelow___spec__1___closed__15; size_t lean_ptr_addr(lean_object*); @@ -185,6 +183,7 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecls___at_Lean_Meta_IndPredBelow_mkCtorType_addMotives___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); +lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_IndPredBelow_backwardsChaining___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelow___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_Meta_IndPredBelow_mkContext___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -284,9 +283,9 @@ lean_object* l_Lean_Expr_constName_x21(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkCtorType_addMotives___spec__1___lambda__1(lean_object*, 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_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_getBelowIndices_loop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkCasesOn___at_Lean_Meta_IndPredBelow_mkBelow___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkInductiveType(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_Meta_IndPredBelow_mkCtorType_addHeaderVars___spec__3(size_t, size_t, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); @@ -302,6 +301,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_IndPredBelow_mkBelow__ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_proveBrecOn_applyCtors___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_checkCount___lambda__2(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*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_IndPredBelow_mkCtorType_mkBelowBinder___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_findBelowIdx___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -335,6 +335,7 @@ static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelo static lean_object* l_Lean_mkCasesOn___at_Lean_Meta_IndPredBelow_mkBelow___spec__1___closed__12; lean_object* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); +lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_1063____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__1; static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7640____closed__10; static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__7; @@ -356,7 +357,6 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Meta_IndPredBelow_proveBrecOn___s static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7640____closed__13; LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkCtorType_mkBelowBinder___spec__3___lambda__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_proveBrecOn_closeGoal___closed__1; -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__2___boxed(lean_object**); lean_object* l_Lean_Meta_introNCore(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_IndPredBelow_mkBelow___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkContext_mkHeader___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -424,7 +424,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBrecOnDecl_mkIH___lambda__2( lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_findBelowIdx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkCtorType_addMotives___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_foldrMUnsafe_fold___at_Lean_Meta_IndPredBelow_mkContext_addMotives___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instHashable; @@ -449,7 +448,6 @@ lean_object* l_Array_ofSubarray___rarg(lean_object*); lean_object* l_Lean_MVarId_apply(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Meta_IndPredBelow_backwardsChaining___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_mkBelowMatcher___spec__3___lambda__2___closed__1; -lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___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*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__5; @@ -463,6 +461,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_IndPredBelow_mkCtorTyp size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_mkCtorType_addHeaderVars___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_proveBrecOn_introNPRec(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_Meta_IndPredBelow_mkCtorType_replaceTempVars___spec__1(lean_object*, size_t, size_t, lean_object*); @@ -532,7 +531,6 @@ lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_ob static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_IndPredBelow_proveBrecOn_applyIH___spec__2___closed__1; static lean_object* l_Lean_Meta_IndPredBelow_proveBrecOn_closeGoal___closed__4; lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___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_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_mkBrecOnDecl_mkType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_IndPredBelow_mkBelowMatcher_convertToBelow___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Meta_IndPredBelow_proveBrecOn_applyCtors___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -570,6 +568,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_IndPredBelow_mkBrecOnDecl_mkIH(lean_object* lean_object* l_List_zipWith___at_List_zip___spec__1___rarg(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_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__16___boxed(lean_object*, lean_object*); +uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_IR_IRType_beq___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_proveBrecOn_applyCtors___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_Meta_transform_visit_visitLet___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__6(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Meta_IndPredBelow_mkCtorType_checkCount___spec__9___boxed(lean_object*, lean_object*); @@ -683,7 +682,7 @@ lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__2; x_3 = l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__5; x_4 = l_Lean_Meta_IndPredBelow_initFn____x40_Lean_Meta_IndPredBelow___hyg_7____closed__9; -x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_RecDepth___hyg_6____spec__1(x_2, x_3, x_4, x_1); +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_Options___hyg_1063____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -15573,7 +15572,7 @@ lean_closure_set(x_20, 1, x_1); lean_closure_set(x_20, 2, x_3); lean_closure_set(x_20, 3, x_4); x_21 = l_Lean_Meta_IndPredBelow_getBelowIndices_loop___closed__1; -x_22 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_2, x_21, x_20, x_6, x_7, x_8, x_9, x_19); +x_22 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_2, x_21, x_20, x_6, x_7, x_8, x_9, x_19); return x_22; } else @@ -17221,44 +17220,7 @@ x_11 = l_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop(x_1, x_3, x_ return x_11; } } -LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___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; -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; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 0); -x_7 = lean_ctor_get(x_2, 0); -x_8 = lean_name_eq(x_6, x_7); -return x_8; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { _start: { if (lean_obj_tag(x_10) == 5) @@ -17288,7 +17250,7 @@ lean_dec(x_10); lean_inc(x_2); x_25 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_25, 0, x_2); -x_26 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(x_24, x_25); +x_26 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_IR_IRType_beq___spec__1(x_24, x_25); lean_dec(x_25); lean_dec(x_24); if (x_26 == 0) @@ -17677,7 +17639,7 @@ x_44 = lean_mk_array(x_42, x_43); x_45 = lean_unsigned_to_nat(1u); x_46 = lean_nat_sub(x_42, x_45); lean_dec(x_42); -x_47 = l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__2(x_1, x_2, x_4, x_5, x_6, x_31, x_32, x_35, x_37, x_39, x_44, x_46, x_7, x_8, x_9, x_10, x_40); +x_47 = l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(x_1, x_2, x_4, x_5, x_6, x_31, x_32, x_35, x_37, x_39, x_44, x_46, x_7, x_8, x_9, x_10, x_40); return x_47; } else @@ -17807,18 +17769,7 @@ lean_dec(x_3); return x_9; } } -LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___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_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___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_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__2___boxed(lean_object** _args) { +LEAN_EXPORT lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1___boxed(lean_object** _args) { lean_object* x_1 = _args[0]; lean_object* x_2 = _args[1]; lean_object* x_3 = _args[2]; @@ -17839,7 +17790,7 @@ lean_object* x_17 = _args[16]; _start: { lean_object* x_18; -x_18 = l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +x_18 = l_Lean_Expr_withAppAux___at_Lean_Meta_IndPredBelow_mkBelowMatcher_transformFields_loop___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16, x_17); return x_18; } } @@ -18307,7 +18258,7 @@ lean_closure_set(x_10, 0, x_1); lean_closure_set(x_10, 1, x_2); lean_closure_set(x_10, 2, x_3); x_11 = 0; -x_12 = l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(x_9, x_10, x_11, x_4, x_5, x_6, x_7, x_8); +x_12 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg(x_9, x_10, x_11, x_4, x_5, x_6, x_7, x_8); return x_12; } } @@ -18977,7 +18928,7 @@ lean_closure_set(x_26, 2, x_6); lean_closure_set(x_26, 3, x_20); x_27 = 0; x_28 = lean_box(x_27); -x_29 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg___boxed), 8, 3); +x_29 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed), 8, 3); lean_closure_set(x_29, 0, x_20); lean_closure_set(x_29, 1, x_26); lean_closure_set(x_29, 2, x_28); @@ -19760,7 +19711,7 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_17 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_13, x_15, x_16, x_5, x_6, x_7, x_8, x_12); +x_17 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_13, x_15, x_16, x_5, x_6, x_7, x_8, x_12); if (lean_obj_tag(x_17) == 0) { lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; diff --git a/stage0/stdlib/Lean/Meta/Injective.c b/stage0/stdlib/Lean/Meta/Injective.c index d9cf4544b68c..92d682b4e29d 100644 --- a/stage0/stdlib/Lean/Meta/Injective.c +++ b/stage0/stdlib/Lean/Meta/Injective.c @@ -13,112 +13,125 @@ #ifdef __cplusplus extern "C" { #endif -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; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__4; lean_object* l_Lean_Meta_mkEqHEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_assumptionCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_mkRecursorInfoForKernelRec___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f(lean_object*); 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*); +static double l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__4; +extern lean_object* l_Lean_trace_profiler_useHearbeats; 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*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__9; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___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_genInjectivity; +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_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_EXPORT lean_object* l_Lean_Meta_occursOrInType___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_Injective___hyg_2297____closed__7; 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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__1; 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*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___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* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__10; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkInjectiveTheorems___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__11; uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__4; -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*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_occursOrInType___lambda__1(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*); +lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__1; 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; +uint8_t l_Lean_HashSetImp_contains___at_Lean_Expr_FindImpl_checkVisited___spec__8(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*); lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27___spec__1___rarg(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__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__5; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkInjectiveTheorems___spec__2___closed__2; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958_(lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremType_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_refl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__5; +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheorem___closed__1; -lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_get_num_heartbeats(lean_object*); +static double l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__5; lean_object* l_Lean_stringToMessageData(lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkInjectiveTheorems___spec__4(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__5___closed__1; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -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_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_MVarId_splitAndCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__7; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10___closed__2; 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*); +lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__4; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__5; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__9; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__5(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*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___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_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__2; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___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*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__1; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__1; static lean_object* l_Lean_Meta_elimOptParam___lambda__1___closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___rarg(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__9(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(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_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 double l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__6; +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__12; +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1; 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*); @@ -127,59 +140,60 @@ 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; lean_object* l_Lean_Meta_isInductivePredicate(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_elimOptParam___lambda__1___closed__3; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__4; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___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*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkInjectiveTheorems___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); -static lean_object* l_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*); +static lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10___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*); +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__3; lean_object* l_Lean_MVarId_casesAnd(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__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremValue___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__3; 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_Name_str___override(lean_object*, lean_object*); lean_object* l_Lean_Meta_injection(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_elimOptParam___lambda__1___closed__1; static lean_object* l_Lean_Meta_mkInjectiveTheoremNameFor___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Meta_mkInjectiveTheorems___spec__1___boxed(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_occursOrInType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__13; extern lean_object* l_Lean_Meta_simpExtension; +extern lean_object* l_Lean_trace_profiler_threshold; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremValue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__6; +LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, 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* 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; @@ -188,13 +202,15 @@ lean_object* l_Lean_Meta_addSimpTheorem(lean_object*, lean_object*, uint8_t, uin double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); 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; -static double l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___spec__8___closed__1; +static double l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__1; lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_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*, 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_mkInjectiveTheorems___spec__5___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___spec__1(lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__4; @@ -203,47 +219,46 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__2; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297____closed__8; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_HashSetImp_insert___at_Lean_Expr_FindImpl_checkVisited___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Meta_mkInjectiveTheorems___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkInjectiveTheorems___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkInjectiveTheoremNameFor___closed__2; static lean_object* l_panic___at___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___spec__1___closed__1; lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1(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___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_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); -lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__6; lean_object* l_Lean_mkArrow(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__8(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_mkInjectiveTheorems___spec__5___lambda__4___closed__2; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__12; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_apply(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__13; static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveTheoremNameFor(lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_mkInjectiveTheorems___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_occurs(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_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f___lambda__3(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam___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_Meta_initFn____x40_Lean_Meta_Injective___hyg_2297_(lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkInjectiveEqTheoremNameFor(lean_object*); static lean_object* l_Lean_Meta_mkInjectiveTheorems___closed__4; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__1; 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_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_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_occursOrInType___lambda__2(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_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_object* lean_array_get_size(lean_object*); @@ -252,32 +267,35 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Injective_0_ static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_mkInjectiveTheorems___spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkInjectiveTheorems___spec__2___boxed(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_List_forIn_loop___at_Lean_Meta_mkInjectiveTheorems___spec__10___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* l_Lean_Meta_intro1Core(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_mkInjectiveTheorems___spec__3___boxed(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__1; 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*); +static lean_object* l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___spec__1___closed__1; uint8_t l_Lean_Exception_isRuntime(lean_object*); lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Expr_isFVar(lean_object*); lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_injTheoremFailureHeader___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__7; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_expr_instantiate1(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__2; static lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveEqTheoremValue___lambda__2___closed__5; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___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_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_elimOptParam(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_withNewBinderInfos___at_Lean_Meta_withInstImplicitAsImplict___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_2157____closed__11; double lean_float_sub(double, double); uint8_t l_Array_isEmpty___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Subarray_forInUnsafe_loop___at___private_Lean_Meta_Injective_0__Lean_Meta_mkAnd_x3f___spec__1___closed__1() { _start: { @@ -503,6 +521,3069 @@ lean_dec(x_2); return x_5; } } +static lean_object* _init_l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___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) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = l_Lean_HashSetImp_insert___at_Lean_Expr_FindImpl_checkVisited___spec__1(x_3, x_1); +x_10 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1___closed__1; +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_8); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; uint8_t x_473; +lean_inc(x_2); +lean_inc(x_3); +x_473 = l_Lean_HashSetImp_contains___at_Lean_Expr_FindImpl_checkVisited___spec__8(x_3, x_2); +if (x_473 == 0) +{ +lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; +x_474 = lean_box(0); +lean_inc(x_2); +x_475 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___lambda__1(x_2, x_474, x_3, x_4, x_5, x_6, x_7, x_8); +x_476 = lean_ctor_get(x_475, 0); +lean_inc(x_476); +x_477 = lean_ctor_get(x_475, 1); +lean_inc(x_477); +lean_dec(x_475); +x_9 = x_476; +x_10 = x_477; +goto block_472; +} +else +{ +lean_object* x_478; lean_object* x_479; +x_478 = lean_box(0); +x_479 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_479, 0, x_478); +lean_ctor_set(x_479, 1, x_3); +x_9 = x_479; +x_10 = x_8; +goto block_472; +} +block_472: +{ +lean_object* x_11; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +if (lean_obj_tag(x_11) == 0) +{ +uint8_t x_12; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_9, 0); +lean_dec(x_13); +x_14 = lean_box(0); +lean_ctor_set(x_9, 0, x_14); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_9); +lean_ctor_set(x_15, 1, x_10); +return x_15; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_9, 1); +lean_inc(x_16); +lean_dec(x_9); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, 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_10); +return x_19; +} +} +else +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_11); +if (x_20 == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_11, 0); +lean_dec(x_21); +x_22 = !lean_is_exclusive(x_9); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_9, 1); +x_24 = lean_ctor_get(x_9, 0); +lean_dec(x_24); +lean_inc(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +x_25 = lean_apply_6(x_1, x_2, x_4, x_5, x_6, x_7, x_10); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) +{ +lean_free_object(x_11); +switch (lean_obj_tag(x_2)) { +case 5: +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_free_object(x_9); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_ctor_get(x_2, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 1); +lean_inc(x_30); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_31 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_29, x_23, x_4, x_5, x_6, x_7, x_28); +if (lean_obj_tag(x_31) == 0) +{ +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +if (lean_obj_tag(x_33) == 0) +{ +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_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_2 = x_30; +x_3 = x_35; +x_8 = x_34; +goto _start; +} +else +{ +uint8_t x_37; +lean_dec(x_30); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_37 = !lean_is_exclusive(x_31); +if (x_37 == 0) +{ +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_31, 0); +lean_dec(x_38); +x_39 = !lean_is_exclusive(x_32); +if (x_39 == 0) +{ +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_32, 0); +lean_dec(x_40); +x_41 = !lean_is_exclusive(x_33); +if (x_41 == 0) +{ +return x_31; +} +else +{ +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_33, 0); +lean_inc(x_42); +lean_dec(x_33); +x_43 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_32, 0, x_43); +return x_31; +} +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_44 = lean_ctor_get(x_32, 1); +lean_inc(x_44); +lean_dec(x_32); +x_45 = lean_ctor_get(x_33, 0); +lean_inc(x_45); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + x_46 = x_33; +} else { + lean_dec_ref(x_33); + x_46 = lean_box(0); +} +if (lean_is_scalar(x_46)) { + x_47 = lean_alloc_ctor(1, 1, 0); +} else { + x_47 = x_46; +} +lean_ctor_set(x_47, 0, x_45); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_44); +lean_ctor_set(x_31, 0, x_48); +return x_31; +} +} +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; +x_49 = lean_ctor_get(x_31, 1); +lean_inc(x_49); +lean_dec(x_31); +x_50 = lean_ctor_get(x_32, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_32)) { + lean_ctor_release(x_32, 0); + lean_ctor_release(x_32, 1); + x_51 = x_32; +} else { + lean_dec_ref(x_32); + x_51 = lean_box(0); +} +x_52 = lean_ctor_get(x_33, 0); +lean_inc(x_52); +if (lean_is_exclusive(x_33)) { + lean_ctor_release(x_33, 0); + x_53 = x_33; +} else { + lean_dec_ref(x_33); + x_53 = lean_box(0); +} +if (lean_is_scalar(x_53)) { + x_54 = lean_alloc_ctor(1, 1, 0); +} else { + x_54 = x_53; +} +lean_ctor_set(x_54, 0, x_52); +if (lean_is_scalar(x_51)) { + x_55 = lean_alloc_ctor(0, 2, 0); +} else { + x_55 = x_51; +} +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_50); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_49); +return x_56; +} +} +} +else +{ +uint8_t x_57; +lean_dec(x_30); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_31); +if (x_57 == 0) +{ +return x_31; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_31, 0); +x_59 = lean_ctor_get(x_31, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_31); +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; +} +} +} +case 6: +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +lean_free_object(x_9); +x_61 = lean_ctor_get(x_25, 1); +lean_inc(x_61); +lean_dec(x_25); +x_62 = lean_ctor_get(x_2, 1); +lean_inc(x_62); +x_63 = lean_ctor_get(x_2, 2); +lean_inc(x_63); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_64 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_62, x_23, x_4, x_5, x_6, x_7, x_61); +if (lean_obj_tag(x_64) == 0) +{ +lean_object* x_65; lean_object* x_66; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_64, 1); +lean_inc(x_67); +lean_dec(x_64); +x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_68); +lean_dec(x_65); +x_2 = x_63; +x_3 = x_68; +x_8 = x_67; +goto _start; +} +else +{ +uint8_t x_70; +lean_dec(x_63); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_64); +if (x_70 == 0) +{ +lean_object* x_71; uint8_t x_72; +x_71 = lean_ctor_get(x_64, 0); +lean_dec(x_71); +x_72 = !lean_is_exclusive(x_65); +if (x_72 == 0) +{ +lean_object* x_73; uint8_t x_74; +x_73 = lean_ctor_get(x_65, 0); +lean_dec(x_73); +x_74 = !lean_is_exclusive(x_66); +if (x_74 == 0) +{ +return x_64; +} +else +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_66, 0); +lean_inc(x_75); +lean_dec(x_66); +x_76 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_65, 0, x_76); +return x_64; +} +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_77 = lean_ctor_get(x_65, 1); +lean_inc(x_77); +lean_dec(x_65); +x_78 = lean_ctor_get(x_66, 0); +lean_inc(x_78); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + x_79 = x_66; +} else { + lean_dec_ref(x_66); + x_79 = lean_box(0); +} +if (lean_is_scalar(x_79)) { + x_80 = lean_alloc_ctor(1, 1, 0); +} else { + x_80 = x_79; +} +lean_ctor_set(x_80, 0, x_78); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_77); +lean_ctor_set(x_64, 0, x_81); +return x_64; +} +} +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; +x_82 = lean_ctor_get(x_64, 1); +lean_inc(x_82); +lean_dec(x_64); +x_83 = lean_ctor_get(x_65, 1); +lean_inc(x_83); +if (lean_is_exclusive(x_65)) { + lean_ctor_release(x_65, 0); + lean_ctor_release(x_65, 1); + x_84 = x_65; +} else { + lean_dec_ref(x_65); + x_84 = lean_box(0); +} +x_85 = lean_ctor_get(x_66, 0); +lean_inc(x_85); +if (lean_is_exclusive(x_66)) { + lean_ctor_release(x_66, 0); + x_86 = x_66; +} else { + lean_dec_ref(x_66); + x_86 = lean_box(0); +} +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 1, 0); +} else { + x_87 = x_86; +} +lean_ctor_set(x_87, 0, x_85); +if (lean_is_scalar(x_84)) { + x_88 = lean_alloc_ctor(0, 2, 0); +} else { + x_88 = x_84; +} +lean_ctor_set(x_88, 0, x_87); +lean_ctor_set(x_88, 1, x_83); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_82); +return x_89; +} +} +} +else +{ +uint8_t x_90; +lean_dec(x_63); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_90 = !lean_is_exclusive(x_64); +if (x_90 == 0) +{ +return x_64; +} +else +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_64, 0); +x_92 = lean_ctor_get(x_64, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_64); +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; +} +} +} +case 7: +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; +lean_free_object(x_9); +x_94 = lean_ctor_get(x_25, 1); +lean_inc(x_94); +lean_dec(x_25); +x_95 = lean_ctor_get(x_2, 1); +lean_inc(x_95); +x_96 = lean_ctor_get(x_2, 2); +lean_inc(x_96); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_97 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_95, x_23, x_4, x_5, x_6, x_7, x_94); +if (lean_obj_tag(x_97) == 0) +{ +lean_object* x_98; lean_object* x_99; +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +x_101 = lean_ctor_get(x_98, 1); +lean_inc(x_101); +lean_dec(x_98); +x_2 = x_96; +x_3 = x_101; +x_8 = x_100; +goto _start; +} +else +{ +uint8_t x_103; +lean_dec(x_96); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_103 = !lean_is_exclusive(x_97); +if (x_103 == 0) +{ +lean_object* x_104; uint8_t x_105; +x_104 = lean_ctor_get(x_97, 0); +lean_dec(x_104); +x_105 = !lean_is_exclusive(x_98); +if (x_105 == 0) +{ +lean_object* x_106; uint8_t x_107; +x_106 = lean_ctor_get(x_98, 0); +lean_dec(x_106); +x_107 = !lean_is_exclusive(x_99); +if (x_107 == 0) +{ +return x_97; +} +else +{ +lean_object* x_108; lean_object* x_109; +x_108 = lean_ctor_get(x_99, 0); +lean_inc(x_108); +lean_dec(x_99); +x_109 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_98, 0, x_109); +return x_97; +} +} +else +{ +lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_110 = lean_ctor_get(x_98, 1); +lean_inc(x_110); +lean_dec(x_98); +x_111 = lean_ctor_get(x_99, 0); +lean_inc(x_111); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + x_112 = x_99; +} else { + lean_dec_ref(x_99); + x_112 = lean_box(0); +} +if (lean_is_scalar(x_112)) { + x_113 = lean_alloc_ctor(1, 1, 0); +} else { + x_113 = x_112; +} +lean_ctor_set(x_113, 0, x_111); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_110); +lean_ctor_set(x_97, 0, x_114); +return x_97; +} +} +else +{ +lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_115 = lean_ctor_get(x_97, 1); +lean_inc(x_115); +lean_dec(x_97); +x_116 = lean_ctor_get(x_98, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_117 = x_98; +} else { + lean_dec_ref(x_98); + x_117 = lean_box(0); +} +x_118 = lean_ctor_get(x_99, 0); +lean_inc(x_118); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + x_119 = x_99; +} else { + lean_dec_ref(x_99); + x_119 = lean_box(0); +} +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 1, 0); +} else { + x_120 = x_119; +} +lean_ctor_set(x_120, 0, x_118); +if (lean_is_scalar(x_117)) { + x_121 = lean_alloc_ctor(0, 2, 0); +} else { + x_121 = x_117; +} +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_116); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_121); +lean_ctor_set(x_122, 1, x_115); +return x_122; +} +} +} +else +{ +uint8_t x_123; +lean_dec(x_96); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_123 = !lean_is_exclusive(x_97); +if (x_123 == 0) +{ +return x_97; +} +else +{ +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_97, 0); +x_125 = lean_ctor_get(x_97, 1); +lean_inc(x_125); +lean_inc(x_124); +lean_dec(x_97); +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 8: +{ +lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; +lean_free_object(x_9); +x_127 = lean_ctor_get(x_25, 1); +lean_inc(x_127); +lean_dec(x_25); +x_128 = lean_ctor_get(x_2, 1); +lean_inc(x_128); +x_129 = lean_ctor_get(x_2, 2); +lean_inc(x_129); +x_130 = lean_ctor_get(x_2, 3); +lean_inc(x_130); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_131 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_128, x_23, x_4, x_5, x_6, x_7, x_127); +if (lean_obj_tag(x_131) == 0) +{ +lean_object* x_132; lean_object* x_133; +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_132, 0); +lean_inc(x_133); +if (lean_obj_tag(x_133) == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_134 = lean_ctor_get(x_131, 1); +lean_inc(x_134); +lean_dec(x_131); +x_135 = lean_ctor_get(x_132, 1); +lean_inc(x_135); +lean_dec(x_132); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_136 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_129, x_135, x_4, x_5, x_6, x_7, x_134); +if (lean_obj_tag(x_136) == 0) +{ +lean_object* x_137; lean_object* x_138; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +if (lean_obj_tag(x_138) == 0) +{ +lean_object* x_139; lean_object* x_140; +x_139 = lean_ctor_get(x_136, 1); +lean_inc(x_139); +lean_dec(x_136); +x_140 = lean_ctor_get(x_137, 1); +lean_inc(x_140); +lean_dec(x_137); +x_2 = x_130; +x_3 = x_140; +x_8 = x_139; +goto _start; +} +else +{ +uint8_t x_142; +lean_dec(x_130); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_142 = !lean_is_exclusive(x_136); +if (x_142 == 0) +{ +lean_object* x_143; uint8_t x_144; +x_143 = lean_ctor_get(x_136, 0); +lean_dec(x_143); +x_144 = !lean_is_exclusive(x_137); +if (x_144 == 0) +{ +lean_object* x_145; uint8_t x_146; +x_145 = lean_ctor_get(x_137, 0); +lean_dec(x_145); +x_146 = !lean_is_exclusive(x_138); +if (x_146 == 0) +{ +return x_136; +} +else +{ +lean_object* x_147; lean_object* x_148; +x_147 = lean_ctor_get(x_138, 0); +lean_inc(x_147); +lean_dec(x_138); +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_137, 0, x_148); +return x_136; +} +} +else +{ +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; +x_149 = lean_ctor_get(x_137, 1); +lean_inc(x_149); +lean_dec(x_137); +x_150 = lean_ctor_get(x_138, 0); +lean_inc(x_150); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + x_151 = x_138; +} else { + lean_dec_ref(x_138); + x_151 = lean_box(0); +} +if (lean_is_scalar(x_151)) { + x_152 = lean_alloc_ctor(1, 1, 0); +} else { + x_152 = x_151; +} +lean_ctor_set(x_152, 0, x_150); +x_153 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_149); +lean_ctor_set(x_136, 0, x_153); +return x_136; +} +} +else +{ +lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_154 = lean_ctor_get(x_136, 1); +lean_inc(x_154); +lean_dec(x_136); +x_155 = lean_ctor_get(x_137, 1); +lean_inc(x_155); +if (lean_is_exclusive(x_137)) { + lean_ctor_release(x_137, 0); + lean_ctor_release(x_137, 1); + x_156 = x_137; +} else { + lean_dec_ref(x_137); + x_156 = lean_box(0); +} +x_157 = lean_ctor_get(x_138, 0); +lean_inc(x_157); +if (lean_is_exclusive(x_138)) { + lean_ctor_release(x_138, 0); + x_158 = x_138; +} else { + lean_dec_ref(x_138); + x_158 = lean_box(0); +} +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(1, 1, 0); +} else { + x_159 = x_158; +} +lean_ctor_set(x_159, 0, x_157); +if (lean_is_scalar(x_156)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_156; +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_155); +x_161 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_161, 0, x_160); +lean_ctor_set(x_161, 1, x_154); +return x_161; +} +} +} +else +{ +uint8_t x_162; +lean_dec(x_130); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_162 = !lean_is_exclusive(x_136); +if (x_162 == 0) +{ +return x_136; +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; +x_163 = lean_ctor_get(x_136, 0); +x_164 = lean_ctor_get(x_136, 1); +lean_inc(x_164); +lean_inc(x_163); +lean_dec(x_136); +x_165 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_165, 0, x_163); +lean_ctor_set(x_165, 1, x_164); +return x_165; +} +} +} +else +{ +uint8_t x_166; +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_166 = !lean_is_exclusive(x_131); +if (x_166 == 0) +{ +lean_object* x_167; uint8_t x_168; +x_167 = lean_ctor_get(x_131, 0); +lean_dec(x_167); +x_168 = !lean_is_exclusive(x_132); +if (x_168 == 0) +{ +lean_object* x_169; uint8_t x_170; +x_169 = lean_ctor_get(x_132, 0); +lean_dec(x_169); +x_170 = !lean_is_exclusive(x_133); +if (x_170 == 0) +{ +return x_131; +} +else +{ +lean_object* x_171; lean_object* x_172; +x_171 = lean_ctor_get(x_133, 0); +lean_inc(x_171); +lean_dec(x_133); +x_172 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_172, 0, x_171); +lean_ctor_set(x_132, 0, x_172); +return x_131; +} +} +else +{ +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_173 = lean_ctor_get(x_132, 1); +lean_inc(x_173); +lean_dec(x_132); +x_174 = lean_ctor_get(x_133, 0); +lean_inc(x_174); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + x_175 = x_133; +} else { + lean_dec_ref(x_133); + x_175 = lean_box(0); +} +if (lean_is_scalar(x_175)) { + x_176 = lean_alloc_ctor(1, 1, 0); +} else { + x_176 = x_175; +} +lean_ctor_set(x_176, 0, x_174); +x_177 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_177, 0, x_176); +lean_ctor_set(x_177, 1, x_173); +lean_ctor_set(x_131, 0, x_177); +return x_131; +} +} +else +{ +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_178 = lean_ctor_get(x_131, 1); +lean_inc(x_178); +lean_dec(x_131); +x_179 = lean_ctor_get(x_132, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_180 = x_132; +} else { + lean_dec_ref(x_132); + x_180 = lean_box(0); +} +x_181 = lean_ctor_get(x_133, 0); +lean_inc(x_181); +if (lean_is_exclusive(x_133)) { + lean_ctor_release(x_133, 0); + x_182 = x_133; +} else { + lean_dec_ref(x_133); + x_182 = lean_box(0); +} +if (lean_is_scalar(x_182)) { + x_183 = lean_alloc_ctor(1, 1, 0); +} else { + x_183 = x_182; +} +lean_ctor_set(x_183, 0, x_181); +if (lean_is_scalar(x_180)) { + x_184 = lean_alloc_ctor(0, 2, 0); +} else { + x_184 = x_180; +} +lean_ctor_set(x_184, 0, x_183); +lean_ctor_set(x_184, 1, x_179); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_178); +return x_185; +} +} +} +else +{ +uint8_t x_186; +lean_dec(x_130); +lean_dec(x_129); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_186 = !lean_is_exclusive(x_131); +if (x_186 == 0) +{ +return x_131; +} +else +{ +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_131, 0); +x_188 = lean_ctor_get(x_131, 1); +lean_inc(x_188); +lean_inc(x_187); +lean_dec(x_131); +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 10: +{ +lean_object* x_190; lean_object* x_191; +lean_free_object(x_9); +x_190 = lean_ctor_get(x_25, 1); +lean_inc(x_190); +lean_dec(x_25); +x_191 = lean_ctor_get(x_2, 1); +lean_inc(x_191); +lean_dec(x_2); +x_2 = x_191; +x_3 = x_23; +x_8 = x_190; +goto _start; +} +case 11: +{ +lean_object* x_193; lean_object* x_194; +lean_free_object(x_9); +x_193 = lean_ctor_get(x_25, 1); +lean_inc(x_193); +lean_dec(x_25); +x_194 = lean_ctor_get(x_2, 2); +lean_inc(x_194); +lean_dec(x_2); +x_2 = x_194; +x_3 = x_23; +x_8 = x_193; +goto _start; +} +default: +{ +uint8_t x_196; +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_196 = !lean_is_exclusive(x_25); +if (x_196 == 0) +{ +lean_object* x_197; lean_object* x_198; +x_197 = lean_ctor_get(x_25, 0); +lean_dec(x_197); +x_198 = lean_box(0); +lean_ctor_set(x_9, 0, x_198); +lean_ctor_set(x_25, 0, x_9); +return x_25; +} +else +{ +lean_object* x_199; lean_object* x_200; lean_object* x_201; +x_199 = lean_ctor_get(x_25, 1); +lean_inc(x_199); +lean_dec(x_25); +x_200 = lean_box(0); +lean_ctor_set(x_9, 0, x_200); +x_201 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_201, 0, x_9); +lean_ctor_set(x_201, 1, x_199); +return x_201; +} +} +} +} +else +{ +uint8_t x_202; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_202 = !lean_is_exclusive(x_25); +if (x_202 == 0) +{ +lean_object* x_203; +x_203 = lean_ctor_get(x_25, 0); +lean_dec(x_203); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_25, 0, x_9); +return x_25; +} +else +{ +lean_object* x_204; lean_object* x_205; +x_204 = lean_ctor_get(x_25, 1); +lean_inc(x_204); +lean_dec(x_25); +lean_ctor_set(x_11, 0, x_2); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_9); +lean_ctor_set(x_205, 1, x_204); +return x_205; +} +} +} +else +{ +uint8_t x_206; +lean_free_object(x_9); +lean_dec(x_23); +lean_free_object(x_11); +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_206 = !lean_is_exclusive(x_25); +if (x_206 == 0) +{ +return x_25; +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; +x_207 = lean_ctor_get(x_25, 0); +x_208 = lean_ctor_get(x_25, 1); +lean_inc(x_208); +lean_inc(x_207); +lean_dec(x_25); +x_209 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); +return x_209; +} +} +} +else +{ +lean_object* x_210; lean_object* x_211; +x_210 = lean_ctor_get(x_9, 1); +lean_inc(x_210); +lean_dec(x_9); +lean_inc(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +x_211 = lean_apply_6(x_1, x_2, x_4, x_5, x_6, x_7, x_10); +if (lean_obj_tag(x_211) == 0) +{ +lean_object* x_212; uint8_t x_213; +x_212 = lean_ctor_get(x_211, 0); +lean_inc(x_212); +x_213 = lean_unbox(x_212); +lean_dec(x_212); +if (x_213 == 0) +{ +lean_free_object(x_11); +switch (lean_obj_tag(x_2)) { +case 5: +{ +lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +x_214 = lean_ctor_get(x_211, 1); +lean_inc(x_214); +lean_dec(x_211); +x_215 = lean_ctor_get(x_2, 0); +lean_inc(x_215); +x_216 = lean_ctor_get(x_2, 1); +lean_inc(x_216); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_217 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_215, x_210, x_4, x_5, x_6, x_7, x_214); +if (lean_obj_tag(x_217) == 0) +{ +lean_object* x_218; lean_object* x_219; +x_218 = lean_ctor_get(x_217, 0); +lean_inc(x_218); +x_219 = lean_ctor_get(x_218, 0); +lean_inc(x_219); +if (lean_obj_tag(x_219) == 0) +{ +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_217, 1); +lean_inc(x_220); +lean_dec(x_217); +x_221 = lean_ctor_get(x_218, 1); +lean_inc(x_221); +lean_dec(x_218); +x_2 = x_216; +x_3 = x_221; +x_8 = x_220; +goto _start; +} +else +{ +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; +lean_dec(x_216); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_223 = lean_ctor_get(x_217, 1); +lean_inc(x_223); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_224 = x_217; +} else { + lean_dec_ref(x_217); + x_224 = lean_box(0); +} +x_225 = lean_ctor_get(x_218, 1); +lean_inc(x_225); +if (lean_is_exclusive(x_218)) { + lean_ctor_release(x_218, 0); + lean_ctor_release(x_218, 1); + x_226 = x_218; +} else { + lean_dec_ref(x_218); + x_226 = lean_box(0); +} +x_227 = lean_ctor_get(x_219, 0); +lean_inc(x_227); +if (lean_is_exclusive(x_219)) { + lean_ctor_release(x_219, 0); + x_228 = x_219; +} else { + lean_dec_ref(x_219); + x_228 = lean_box(0); +} +if (lean_is_scalar(x_228)) { + x_229 = lean_alloc_ctor(1, 1, 0); +} else { + x_229 = x_228; +} +lean_ctor_set(x_229, 0, x_227); +if (lean_is_scalar(x_226)) { + x_230 = lean_alloc_ctor(0, 2, 0); +} else { + x_230 = x_226; +} +lean_ctor_set(x_230, 0, x_229); +lean_ctor_set(x_230, 1, x_225); +if (lean_is_scalar(x_224)) { + x_231 = lean_alloc_ctor(0, 2, 0); +} else { + x_231 = x_224; +} +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_223); +return x_231; +} +} +else +{ +lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; +lean_dec(x_216); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_232 = lean_ctor_get(x_217, 0); +lean_inc(x_232); +x_233 = lean_ctor_get(x_217, 1); +lean_inc(x_233); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_234 = x_217; +} else { + lean_dec_ref(x_217); + x_234 = lean_box(0); +} +if (lean_is_scalar(x_234)) { + x_235 = lean_alloc_ctor(1, 2, 0); +} else { + x_235 = x_234; +} +lean_ctor_set(x_235, 0, x_232); +lean_ctor_set(x_235, 1, x_233); +return x_235; +} +} +case 6: +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; +x_236 = lean_ctor_get(x_211, 1); +lean_inc(x_236); +lean_dec(x_211); +x_237 = lean_ctor_get(x_2, 1); +lean_inc(x_237); +x_238 = lean_ctor_get(x_2, 2); +lean_inc(x_238); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_239 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_237, x_210, x_4, x_5, x_6, x_7, x_236); +if (lean_obj_tag(x_239) == 0) +{ +lean_object* x_240; lean_object* x_241; +x_240 = lean_ctor_get(x_239, 0); +lean_inc(x_240); +x_241 = lean_ctor_get(x_240, 0); +lean_inc(x_241); +if (lean_obj_tag(x_241) == 0) +{ +lean_object* x_242; lean_object* x_243; +x_242 = lean_ctor_get(x_239, 1); +lean_inc(x_242); +lean_dec(x_239); +x_243 = lean_ctor_get(x_240, 1); +lean_inc(x_243); +lean_dec(x_240); +x_2 = x_238; +x_3 = x_243; +x_8 = x_242; +goto _start; +} +else +{ +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_dec(x_238); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_245 = lean_ctor_get(x_239, 1); +lean_inc(x_245); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_246 = x_239; +} else { + lean_dec_ref(x_239); + x_246 = lean_box(0); +} +x_247 = lean_ctor_get(x_240, 1); +lean_inc(x_247); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_248 = x_240; +} else { + lean_dec_ref(x_240); + x_248 = lean_box(0); +} +x_249 = lean_ctor_get(x_241, 0); +lean_inc(x_249); +if (lean_is_exclusive(x_241)) { + lean_ctor_release(x_241, 0); + x_250 = x_241; +} else { + lean_dec_ref(x_241); + x_250 = lean_box(0); +} +if (lean_is_scalar(x_250)) { + x_251 = lean_alloc_ctor(1, 1, 0); +} else { + x_251 = x_250; +} +lean_ctor_set(x_251, 0, x_249); +if (lean_is_scalar(x_248)) { + x_252 = lean_alloc_ctor(0, 2, 0); +} else { + x_252 = x_248; +} +lean_ctor_set(x_252, 0, x_251); +lean_ctor_set(x_252, 1, x_247); +if (lean_is_scalar(x_246)) { + x_253 = lean_alloc_ctor(0, 2, 0); +} else { + x_253 = x_246; +} +lean_ctor_set(x_253, 0, x_252); +lean_ctor_set(x_253, 1, x_245); +return x_253; +} +} +else +{ +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; +lean_dec(x_238); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_254 = lean_ctor_get(x_239, 0); +lean_inc(x_254); +x_255 = lean_ctor_get(x_239, 1); +lean_inc(x_255); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_256 = x_239; +} else { + lean_dec_ref(x_239); + x_256 = lean_box(0); +} +if (lean_is_scalar(x_256)) { + x_257 = lean_alloc_ctor(1, 2, 0); +} else { + x_257 = x_256; +} +lean_ctor_set(x_257, 0, x_254); +lean_ctor_set(x_257, 1, x_255); +return x_257; +} +} +case 7: +{ +lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_258 = lean_ctor_get(x_211, 1); +lean_inc(x_258); +lean_dec(x_211); +x_259 = lean_ctor_get(x_2, 1); +lean_inc(x_259); +x_260 = lean_ctor_get(x_2, 2); +lean_inc(x_260); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_261 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_259, x_210, x_4, x_5, x_6, x_7, x_258); +if (lean_obj_tag(x_261) == 0) +{ +lean_object* x_262; lean_object* x_263; +x_262 = lean_ctor_get(x_261, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +if (lean_obj_tag(x_263) == 0) +{ +lean_object* x_264; lean_object* x_265; +x_264 = lean_ctor_get(x_261, 1); +lean_inc(x_264); +lean_dec(x_261); +x_265 = lean_ctor_get(x_262, 1); +lean_inc(x_265); +lean_dec(x_262); +x_2 = x_260; +x_3 = x_265; +x_8 = x_264; +goto _start; +} +else +{ +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_dec(x_260); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_267 = lean_ctor_get(x_261, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_261)) { + lean_ctor_release(x_261, 0); + lean_ctor_release(x_261, 1); + x_268 = x_261; +} else { + lean_dec_ref(x_261); + x_268 = lean_box(0); +} +x_269 = lean_ctor_get(x_262, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_270 = x_262; +} else { + lean_dec_ref(x_262); + x_270 = lean_box(0); +} +x_271 = lean_ctor_get(x_263, 0); +lean_inc(x_271); +if (lean_is_exclusive(x_263)) { + lean_ctor_release(x_263, 0); + x_272 = x_263; +} else { + lean_dec_ref(x_263); + x_272 = lean_box(0); +} +if (lean_is_scalar(x_272)) { + x_273 = lean_alloc_ctor(1, 1, 0); +} else { + x_273 = x_272; +} +lean_ctor_set(x_273, 0, x_271); +if (lean_is_scalar(x_270)) { + x_274 = lean_alloc_ctor(0, 2, 0); +} else { + x_274 = x_270; +} +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_269); +if (lean_is_scalar(x_268)) { + x_275 = lean_alloc_ctor(0, 2, 0); +} else { + x_275 = x_268; +} +lean_ctor_set(x_275, 0, x_274); +lean_ctor_set(x_275, 1, x_267); +return x_275; +} +} +else +{ +lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; +lean_dec(x_260); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_276 = lean_ctor_get(x_261, 0); +lean_inc(x_276); +x_277 = lean_ctor_get(x_261, 1); +lean_inc(x_277); +if (lean_is_exclusive(x_261)) { + lean_ctor_release(x_261, 0); + lean_ctor_release(x_261, 1); + x_278 = x_261; +} else { + lean_dec_ref(x_261); + x_278 = lean_box(0); +} +if (lean_is_scalar(x_278)) { + x_279 = lean_alloc_ctor(1, 2, 0); +} else { + x_279 = x_278; +} +lean_ctor_set(x_279, 0, x_276); +lean_ctor_set(x_279, 1, x_277); +return x_279; +} +} +case 8: +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; +x_280 = lean_ctor_get(x_211, 1); +lean_inc(x_280); +lean_dec(x_211); +x_281 = lean_ctor_get(x_2, 1); +lean_inc(x_281); +x_282 = lean_ctor_get(x_2, 2); +lean_inc(x_282); +x_283 = lean_ctor_get(x_2, 3); +lean_inc(x_283); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_284 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_281, x_210, x_4, x_5, x_6, x_7, x_280); +if (lean_obj_tag(x_284) == 0) +{ +lean_object* x_285; lean_object* x_286; +x_285 = lean_ctor_get(x_284, 0); +lean_inc(x_285); +x_286 = lean_ctor_get(x_285, 0); +lean_inc(x_286); +if (lean_obj_tag(x_286) == 0) +{ +lean_object* x_287; lean_object* x_288; lean_object* x_289; +x_287 = lean_ctor_get(x_284, 1); +lean_inc(x_287); +lean_dec(x_284); +x_288 = lean_ctor_get(x_285, 1); +lean_inc(x_288); +lean_dec(x_285); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_289 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_282, x_288, x_4, x_5, x_6, x_7, x_287); +if (lean_obj_tag(x_289) == 0) +{ +lean_object* x_290; lean_object* x_291; +x_290 = lean_ctor_get(x_289, 0); +lean_inc(x_290); +x_291 = lean_ctor_get(x_290, 0); +lean_inc(x_291); +if (lean_obj_tag(x_291) == 0) +{ +lean_object* x_292; lean_object* x_293; +x_292 = lean_ctor_get(x_289, 1); +lean_inc(x_292); +lean_dec(x_289); +x_293 = lean_ctor_get(x_290, 1); +lean_inc(x_293); +lean_dec(x_290); +x_2 = x_283; +x_3 = x_293; +x_8 = x_292; +goto _start; +} +else +{ +lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; +lean_dec(x_283); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_295 = lean_ctor_get(x_289, 1); +lean_inc(x_295); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + x_296 = x_289; +} else { + lean_dec_ref(x_289); + x_296 = lean_box(0); +} +x_297 = lean_ctor_get(x_290, 1); +lean_inc(x_297); +if (lean_is_exclusive(x_290)) { + lean_ctor_release(x_290, 0); + lean_ctor_release(x_290, 1); + x_298 = x_290; +} else { + lean_dec_ref(x_290); + x_298 = lean_box(0); +} +x_299 = lean_ctor_get(x_291, 0); +lean_inc(x_299); +if (lean_is_exclusive(x_291)) { + lean_ctor_release(x_291, 0); + x_300 = x_291; +} else { + lean_dec_ref(x_291); + x_300 = lean_box(0); +} +if (lean_is_scalar(x_300)) { + x_301 = lean_alloc_ctor(1, 1, 0); +} else { + x_301 = x_300; +} +lean_ctor_set(x_301, 0, x_299); +if (lean_is_scalar(x_298)) { + x_302 = lean_alloc_ctor(0, 2, 0); +} else { + x_302 = x_298; +} +lean_ctor_set(x_302, 0, x_301); +lean_ctor_set(x_302, 1, x_297); +if (lean_is_scalar(x_296)) { + x_303 = lean_alloc_ctor(0, 2, 0); +} else { + x_303 = x_296; +} +lean_ctor_set(x_303, 0, x_302); +lean_ctor_set(x_303, 1, x_295); +return x_303; +} +} +else +{ +lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; +lean_dec(x_283); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_304 = lean_ctor_get(x_289, 0); +lean_inc(x_304); +x_305 = lean_ctor_get(x_289, 1); +lean_inc(x_305); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + x_306 = x_289; +} else { + lean_dec_ref(x_289); + x_306 = lean_box(0); +} +if (lean_is_scalar(x_306)) { + x_307 = lean_alloc_ctor(1, 2, 0); +} else { + x_307 = x_306; +} +lean_ctor_set(x_307, 0, x_304); +lean_ctor_set(x_307, 1, x_305); +return x_307; +} +} +else +{ +lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; +lean_dec(x_283); +lean_dec(x_282); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_308 = lean_ctor_get(x_284, 1); +lean_inc(x_308); +if (lean_is_exclusive(x_284)) { + lean_ctor_release(x_284, 0); + lean_ctor_release(x_284, 1); + x_309 = x_284; +} else { + lean_dec_ref(x_284); + x_309 = lean_box(0); +} +x_310 = lean_ctor_get(x_285, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_285)) { + lean_ctor_release(x_285, 0); + lean_ctor_release(x_285, 1); + x_311 = x_285; +} else { + lean_dec_ref(x_285); + x_311 = lean_box(0); +} +x_312 = lean_ctor_get(x_286, 0); +lean_inc(x_312); +if (lean_is_exclusive(x_286)) { + lean_ctor_release(x_286, 0); + x_313 = x_286; +} else { + lean_dec_ref(x_286); + x_313 = lean_box(0); +} +if (lean_is_scalar(x_313)) { + x_314 = lean_alloc_ctor(1, 1, 0); +} else { + x_314 = x_313; +} +lean_ctor_set(x_314, 0, x_312); +if (lean_is_scalar(x_311)) { + x_315 = lean_alloc_ctor(0, 2, 0); +} else { + x_315 = x_311; +} +lean_ctor_set(x_315, 0, x_314); +lean_ctor_set(x_315, 1, x_310); +if (lean_is_scalar(x_309)) { + x_316 = lean_alloc_ctor(0, 2, 0); +} else { + x_316 = x_309; +} +lean_ctor_set(x_316, 0, x_315); +lean_ctor_set(x_316, 1, x_308); +return x_316; +} +} +else +{ +lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; +lean_dec(x_283); +lean_dec(x_282); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_317 = lean_ctor_get(x_284, 0); +lean_inc(x_317); +x_318 = lean_ctor_get(x_284, 1); +lean_inc(x_318); +if (lean_is_exclusive(x_284)) { + lean_ctor_release(x_284, 0); + lean_ctor_release(x_284, 1); + x_319 = x_284; +} else { + lean_dec_ref(x_284); + x_319 = lean_box(0); +} +if (lean_is_scalar(x_319)) { + x_320 = lean_alloc_ctor(1, 2, 0); +} else { + x_320 = x_319; +} +lean_ctor_set(x_320, 0, x_317); +lean_ctor_set(x_320, 1, x_318); +return x_320; +} +} +case 10: +{ +lean_object* x_321; lean_object* x_322; +x_321 = lean_ctor_get(x_211, 1); +lean_inc(x_321); +lean_dec(x_211); +x_322 = lean_ctor_get(x_2, 1); +lean_inc(x_322); +lean_dec(x_2); +x_2 = x_322; +x_3 = x_210; +x_8 = x_321; +goto _start; +} +case 11: +{ +lean_object* x_324; lean_object* x_325; +x_324 = lean_ctor_get(x_211, 1); +lean_inc(x_324); +lean_dec(x_211); +x_325 = lean_ctor_get(x_2, 2); +lean_inc(x_325); +lean_dec(x_2); +x_2 = x_325; +x_3 = x_210; +x_8 = x_324; +goto _start; +} +default: +{ +lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +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_327 = lean_ctor_get(x_211, 1); +lean_inc(x_327); +if (lean_is_exclusive(x_211)) { + lean_ctor_release(x_211, 0); + lean_ctor_release(x_211, 1); + x_328 = x_211; +} else { + lean_dec_ref(x_211); + x_328 = lean_box(0); +} +x_329 = lean_box(0); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_210); +if (lean_is_scalar(x_328)) { + x_331 = lean_alloc_ctor(0, 2, 0); +} else { + x_331 = x_328; +} +lean_ctor_set(x_331, 0, x_330); +lean_ctor_set(x_331, 1, x_327); +return x_331; +} +} +} +else +{ +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_332 = lean_ctor_get(x_211, 1); +lean_inc(x_332); +if (lean_is_exclusive(x_211)) { + lean_ctor_release(x_211, 0); + lean_ctor_release(x_211, 1); + x_333 = x_211; +} else { + lean_dec_ref(x_211); + x_333 = lean_box(0); +} +lean_ctor_set(x_11, 0, x_2); +x_334 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_334, 0, x_11); +lean_ctor_set(x_334, 1, x_210); +if (lean_is_scalar(x_333)) { + x_335 = lean_alloc_ctor(0, 2, 0); +} else { + x_335 = x_333; +} +lean_ctor_set(x_335, 0, x_334); +lean_ctor_set(x_335, 1, x_332); +return x_335; +} +} +else +{ +lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; +lean_dec(x_210); +lean_free_object(x_11); +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_336 = lean_ctor_get(x_211, 0); +lean_inc(x_336); +x_337 = lean_ctor_get(x_211, 1); +lean_inc(x_337); +if (lean_is_exclusive(x_211)) { + lean_ctor_release(x_211, 0); + lean_ctor_release(x_211, 1); + x_338 = x_211; +} else { + lean_dec_ref(x_211); + x_338 = lean_box(0); +} +if (lean_is_scalar(x_338)) { + x_339 = lean_alloc_ctor(1, 2, 0); +} else { + x_339 = x_338; +} +lean_ctor_set(x_339, 0, x_336); +lean_ctor_set(x_339, 1, x_337); +return x_339; +} +} +} +else +{ +lean_object* x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_11); +x_340 = lean_ctor_get(x_9, 1); +lean_inc(x_340); +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + x_341 = x_9; +} else { + lean_dec_ref(x_9); + x_341 = lean_box(0); +} +lean_inc(x_1); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_2); +x_342 = lean_apply_6(x_1, x_2, x_4, x_5, x_6, x_7, x_10); +if (lean_obj_tag(x_342) == 0) +{ +lean_object* x_343; uint8_t x_344; +x_343 = lean_ctor_get(x_342, 0); +lean_inc(x_343); +x_344 = lean_unbox(x_343); +lean_dec(x_343); +if (x_344 == 0) +{ +switch (lean_obj_tag(x_2)) { +case 5: +{ +lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; +lean_dec(x_341); +x_345 = lean_ctor_get(x_342, 1); +lean_inc(x_345); +lean_dec(x_342); +x_346 = lean_ctor_get(x_2, 0); +lean_inc(x_346); +x_347 = lean_ctor_get(x_2, 1); +lean_inc(x_347); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_348 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_346, x_340, x_4, x_5, x_6, x_7, x_345); +if (lean_obj_tag(x_348) == 0) +{ +lean_object* x_349; lean_object* x_350; +x_349 = lean_ctor_get(x_348, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_349, 0); +lean_inc(x_350); +if (lean_obj_tag(x_350) == 0) +{ +lean_object* x_351; lean_object* x_352; +x_351 = lean_ctor_get(x_348, 1); +lean_inc(x_351); +lean_dec(x_348); +x_352 = lean_ctor_get(x_349, 1); +lean_inc(x_352); +lean_dec(x_349); +x_2 = x_347; +x_3 = x_352; +x_8 = x_351; +goto _start; +} +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_dec(x_347); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_354 = lean_ctor_get(x_348, 1); +lean_inc(x_354); +if (lean_is_exclusive(x_348)) { + lean_ctor_release(x_348, 0); + lean_ctor_release(x_348, 1); + x_355 = x_348; +} else { + lean_dec_ref(x_348); + x_355 = lean_box(0); +} +x_356 = lean_ctor_get(x_349, 1); +lean_inc(x_356); +if (lean_is_exclusive(x_349)) { + lean_ctor_release(x_349, 0); + lean_ctor_release(x_349, 1); + x_357 = x_349; +} else { + lean_dec_ref(x_349); + x_357 = lean_box(0); +} +x_358 = lean_ctor_get(x_350, 0); +lean_inc(x_358); +if (lean_is_exclusive(x_350)) { + lean_ctor_release(x_350, 0); + x_359 = x_350; +} else { + lean_dec_ref(x_350); + x_359 = lean_box(0); +} +if (lean_is_scalar(x_359)) { + x_360 = lean_alloc_ctor(1, 1, 0); +} else { + x_360 = x_359; +} +lean_ctor_set(x_360, 0, x_358); +if (lean_is_scalar(x_357)) { + x_361 = lean_alloc_ctor(0, 2, 0); +} else { + x_361 = x_357; +} +lean_ctor_set(x_361, 0, x_360); +lean_ctor_set(x_361, 1, x_356); +if (lean_is_scalar(x_355)) { + x_362 = lean_alloc_ctor(0, 2, 0); +} else { + x_362 = x_355; +} +lean_ctor_set(x_362, 0, x_361); +lean_ctor_set(x_362, 1, x_354); +return x_362; +} +} +else +{ +lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; +lean_dec(x_347); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_363 = lean_ctor_get(x_348, 0); +lean_inc(x_363); +x_364 = lean_ctor_get(x_348, 1); +lean_inc(x_364); +if (lean_is_exclusive(x_348)) { + lean_ctor_release(x_348, 0); + lean_ctor_release(x_348, 1); + x_365 = x_348; +} else { + lean_dec_ref(x_348); + x_365 = lean_box(0); +} +if (lean_is_scalar(x_365)) { + x_366 = lean_alloc_ctor(1, 2, 0); +} else { + x_366 = x_365; +} +lean_ctor_set(x_366, 0, x_363); +lean_ctor_set(x_366, 1, x_364); +return x_366; +} +} +case 6: +{ +lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; +lean_dec(x_341); +x_367 = lean_ctor_get(x_342, 1); +lean_inc(x_367); +lean_dec(x_342); +x_368 = lean_ctor_get(x_2, 1); +lean_inc(x_368); +x_369 = lean_ctor_get(x_2, 2); +lean_inc(x_369); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_370 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_368, x_340, x_4, x_5, x_6, x_7, x_367); +if (lean_obj_tag(x_370) == 0) +{ +lean_object* x_371; lean_object* x_372; +x_371 = lean_ctor_get(x_370, 0); +lean_inc(x_371); +x_372 = lean_ctor_get(x_371, 0); +lean_inc(x_372); +if (lean_obj_tag(x_372) == 0) +{ +lean_object* x_373; lean_object* x_374; +x_373 = lean_ctor_get(x_370, 1); +lean_inc(x_373); +lean_dec(x_370); +x_374 = lean_ctor_get(x_371, 1); +lean_inc(x_374); +lean_dec(x_371); +x_2 = x_369; +x_3 = x_374; +x_8 = x_373; +goto _start; +} +else +{ +lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; +lean_dec(x_369); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_376 = lean_ctor_get(x_370, 1); +lean_inc(x_376); +if (lean_is_exclusive(x_370)) { + lean_ctor_release(x_370, 0); + lean_ctor_release(x_370, 1); + x_377 = x_370; +} else { + lean_dec_ref(x_370); + x_377 = lean_box(0); +} +x_378 = lean_ctor_get(x_371, 1); +lean_inc(x_378); +if (lean_is_exclusive(x_371)) { + lean_ctor_release(x_371, 0); + lean_ctor_release(x_371, 1); + x_379 = x_371; +} else { + lean_dec_ref(x_371); + x_379 = lean_box(0); +} +x_380 = lean_ctor_get(x_372, 0); +lean_inc(x_380); +if (lean_is_exclusive(x_372)) { + lean_ctor_release(x_372, 0); + x_381 = x_372; +} else { + lean_dec_ref(x_372); + x_381 = lean_box(0); +} +if (lean_is_scalar(x_381)) { + x_382 = lean_alloc_ctor(1, 1, 0); +} else { + x_382 = x_381; +} +lean_ctor_set(x_382, 0, x_380); +if (lean_is_scalar(x_379)) { + x_383 = lean_alloc_ctor(0, 2, 0); +} else { + x_383 = x_379; +} +lean_ctor_set(x_383, 0, x_382); +lean_ctor_set(x_383, 1, x_378); +if (lean_is_scalar(x_377)) { + x_384 = lean_alloc_ctor(0, 2, 0); +} else { + x_384 = x_377; +} +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_376); +return x_384; +} +} +else +{ +lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; +lean_dec(x_369); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_385 = lean_ctor_get(x_370, 0); +lean_inc(x_385); +x_386 = lean_ctor_get(x_370, 1); +lean_inc(x_386); +if (lean_is_exclusive(x_370)) { + lean_ctor_release(x_370, 0); + lean_ctor_release(x_370, 1); + x_387 = x_370; +} else { + lean_dec_ref(x_370); + x_387 = lean_box(0); +} +if (lean_is_scalar(x_387)) { + x_388 = lean_alloc_ctor(1, 2, 0); +} else { + x_388 = x_387; +} +lean_ctor_set(x_388, 0, x_385); +lean_ctor_set(x_388, 1, x_386); +return x_388; +} +} +case 7: +{ +lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; +lean_dec(x_341); +x_389 = lean_ctor_get(x_342, 1); +lean_inc(x_389); +lean_dec(x_342); +x_390 = lean_ctor_get(x_2, 1); +lean_inc(x_390); +x_391 = lean_ctor_get(x_2, 2); +lean_inc(x_391); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_392 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_390, x_340, x_4, x_5, x_6, x_7, x_389); +if (lean_obj_tag(x_392) == 0) +{ +lean_object* x_393; lean_object* x_394; +x_393 = lean_ctor_get(x_392, 0); +lean_inc(x_393); +x_394 = lean_ctor_get(x_393, 0); +lean_inc(x_394); +if (lean_obj_tag(x_394) == 0) +{ +lean_object* x_395; lean_object* x_396; +x_395 = lean_ctor_get(x_392, 1); +lean_inc(x_395); +lean_dec(x_392); +x_396 = lean_ctor_get(x_393, 1); +lean_inc(x_396); +lean_dec(x_393); +x_2 = x_391; +x_3 = x_396; +x_8 = x_395; +goto _start; +} +else +{ +lean_object* x_398; lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; +lean_dec(x_391); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_398 = lean_ctor_get(x_392, 1); +lean_inc(x_398); +if (lean_is_exclusive(x_392)) { + lean_ctor_release(x_392, 0); + lean_ctor_release(x_392, 1); + x_399 = x_392; +} else { + lean_dec_ref(x_392); + x_399 = lean_box(0); +} +x_400 = lean_ctor_get(x_393, 1); +lean_inc(x_400); +if (lean_is_exclusive(x_393)) { + lean_ctor_release(x_393, 0); + lean_ctor_release(x_393, 1); + x_401 = x_393; +} else { + lean_dec_ref(x_393); + x_401 = lean_box(0); +} +x_402 = lean_ctor_get(x_394, 0); +lean_inc(x_402); +if (lean_is_exclusive(x_394)) { + lean_ctor_release(x_394, 0); + x_403 = x_394; +} else { + lean_dec_ref(x_394); + x_403 = lean_box(0); +} +if (lean_is_scalar(x_403)) { + x_404 = lean_alloc_ctor(1, 1, 0); +} else { + x_404 = x_403; +} +lean_ctor_set(x_404, 0, x_402); +if (lean_is_scalar(x_401)) { + x_405 = lean_alloc_ctor(0, 2, 0); +} else { + x_405 = x_401; +} +lean_ctor_set(x_405, 0, x_404); +lean_ctor_set(x_405, 1, x_400); +if (lean_is_scalar(x_399)) { + x_406 = lean_alloc_ctor(0, 2, 0); +} else { + x_406 = x_399; +} +lean_ctor_set(x_406, 0, x_405); +lean_ctor_set(x_406, 1, x_398); +return x_406; +} +} +else +{ +lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; +lean_dec(x_391); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_407 = lean_ctor_get(x_392, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_392, 1); +lean_inc(x_408); +if (lean_is_exclusive(x_392)) { + lean_ctor_release(x_392, 0); + lean_ctor_release(x_392, 1); + x_409 = x_392; +} else { + lean_dec_ref(x_392); + x_409 = lean_box(0); +} +if (lean_is_scalar(x_409)) { + x_410 = lean_alloc_ctor(1, 2, 0); +} else { + x_410 = x_409; +} +lean_ctor_set(x_410, 0, x_407); +lean_ctor_set(x_410, 1, x_408); +return x_410; +} +} +case 8: +{ +lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; lean_object* x_415; +lean_dec(x_341); +x_411 = lean_ctor_get(x_342, 1); +lean_inc(x_411); +lean_dec(x_342); +x_412 = lean_ctor_get(x_2, 1); +lean_inc(x_412); +x_413 = lean_ctor_get(x_2, 2); +lean_inc(x_413); +x_414 = lean_ctor_get(x_2, 3); +lean_inc(x_414); +lean_dec(x_2); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_415 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_412, x_340, x_4, x_5, x_6, x_7, x_411); +if (lean_obj_tag(x_415) == 0) +{ +lean_object* x_416; lean_object* x_417; +x_416 = lean_ctor_get(x_415, 0); +lean_inc(x_416); +x_417 = lean_ctor_get(x_416, 0); +lean_inc(x_417); +if (lean_obj_tag(x_417) == 0) +{ +lean_object* x_418; lean_object* x_419; lean_object* x_420; +x_418 = lean_ctor_get(x_415, 1); +lean_inc(x_418); +lean_dec(x_415); +x_419 = lean_ctor_get(x_416, 1); +lean_inc(x_419); +lean_dec(x_416); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_1); +x_420 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_413, x_419, x_4, x_5, x_6, x_7, x_418); +if (lean_obj_tag(x_420) == 0) +{ +lean_object* x_421; lean_object* x_422; +x_421 = lean_ctor_get(x_420, 0); +lean_inc(x_421); +x_422 = lean_ctor_get(x_421, 0); +lean_inc(x_422); +if (lean_obj_tag(x_422) == 0) +{ +lean_object* x_423; lean_object* x_424; +x_423 = lean_ctor_get(x_420, 1); +lean_inc(x_423); +lean_dec(x_420); +x_424 = lean_ctor_get(x_421, 1); +lean_inc(x_424); +lean_dec(x_421); +x_2 = x_414; +x_3 = x_424; +x_8 = x_423; +goto _start; +} +else +{ +lean_object* x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; lean_object* x_432; lean_object* x_433; lean_object* x_434; +lean_dec(x_414); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_426 = lean_ctor_get(x_420, 1); +lean_inc(x_426); +if (lean_is_exclusive(x_420)) { + lean_ctor_release(x_420, 0); + lean_ctor_release(x_420, 1); + x_427 = x_420; +} else { + lean_dec_ref(x_420); + x_427 = lean_box(0); +} +x_428 = lean_ctor_get(x_421, 1); +lean_inc(x_428); +if (lean_is_exclusive(x_421)) { + lean_ctor_release(x_421, 0); + lean_ctor_release(x_421, 1); + x_429 = x_421; +} else { + lean_dec_ref(x_421); + x_429 = lean_box(0); +} +x_430 = lean_ctor_get(x_422, 0); +lean_inc(x_430); +if (lean_is_exclusive(x_422)) { + lean_ctor_release(x_422, 0); + x_431 = x_422; +} else { + lean_dec_ref(x_422); + x_431 = lean_box(0); +} +if (lean_is_scalar(x_431)) { + x_432 = lean_alloc_ctor(1, 1, 0); +} else { + x_432 = x_431; +} +lean_ctor_set(x_432, 0, x_430); +if (lean_is_scalar(x_429)) { + x_433 = lean_alloc_ctor(0, 2, 0); +} else { + x_433 = x_429; +} +lean_ctor_set(x_433, 0, x_432); +lean_ctor_set(x_433, 1, x_428); +if (lean_is_scalar(x_427)) { + x_434 = lean_alloc_ctor(0, 2, 0); +} else { + x_434 = x_427; +} +lean_ctor_set(x_434, 0, x_433); +lean_ctor_set(x_434, 1, x_426); +return x_434; +} +} +else +{ +lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; +lean_dec(x_414); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_435 = lean_ctor_get(x_420, 0); +lean_inc(x_435); +x_436 = lean_ctor_get(x_420, 1); +lean_inc(x_436); +if (lean_is_exclusive(x_420)) { + lean_ctor_release(x_420, 0); + lean_ctor_release(x_420, 1); + x_437 = x_420; +} else { + lean_dec_ref(x_420); + x_437 = lean_box(0); +} +if (lean_is_scalar(x_437)) { + x_438 = lean_alloc_ctor(1, 2, 0); +} else { + x_438 = x_437; +} +lean_ctor_set(x_438, 0, x_435); +lean_ctor_set(x_438, 1, x_436); +return x_438; +} +} +else +{ +lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; lean_object* x_443; lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; +lean_dec(x_414); +lean_dec(x_413); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_439 = lean_ctor_get(x_415, 1); +lean_inc(x_439); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_440 = x_415; +} else { + lean_dec_ref(x_415); + x_440 = lean_box(0); +} +x_441 = lean_ctor_get(x_416, 1); +lean_inc(x_441); +if (lean_is_exclusive(x_416)) { + lean_ctor_release(x_416, 0); + lean_ctor_release(x_416, 1); + x_442 = x_416; +} else { + lean_dec_ref(x_416); + x_442 = lean_box(0); +} +x_443 = lean_ctor_get(x_417, 0); +lean_inc(x_443); +if (lean_is_exclusive(x_417)) { + lean_ctor_release(x_417, 0); + x_444 = x_417; +} else { + lean_dec_ref(x_417); + x_444 = lean_box(0); +} +if (lean_is_scalar(x_444)) { + x_445 = lean_alloc_ctor(1, 1, 0); +} else { + x_445 = x_444; +} +lean_ctor_set(x_445, 0, x_443); +if (lean_is_scalar(x_442)) { + x_446 = lean_alloc_ctor(0, 2, 0); +} else { + x_446 = x_442; +} +lean_ctor_set(x_446, 0, x_445); +lean_ctor_set(x_446, 1, x_441); +if (lean_is_scalar(x_440)) { + x_447 = lean_alloc_ctor(0, 2, 0); +} else { + x_447 = x_440; +} +lean_ctor_set(x_447, 0, x_446); +lean_ctor_set(x_447, 1, x_439); +return x_447; +} +} +else +{ +lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; +lean_dec(x_414); +lean_dec(x_413); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_448 = lean_ctor_get(x_415, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_415, 1); +lean_inc(x_449); +if (lean_is_exclusive(x_415)) { + lean_ctor_release(x_415, 0); + lean_ctor_release(x_415, 1); + x_450 = x_415; +} else { + lean_dec_ref(x_415); + x_450 = lean_box(0); +} +if (lean_is_scalar(x_450)) { + x_451 = lean_alloc_ctor(1, 2, 0); +} else { + x_451 = x_450; +} +lean_ctor_set(x_451, 0, x_448); +lean_ctor_set(x_451, 1, x_449); +return x_451; +} +} +case 10: +{ +lean_object* x_452; lean_object* x_453; +lean_dec(x_341); +x_452 = lean_ctor_get(x_342, 1); +lean_inc(x_452); +lean_dec(x_342); +x_453 = lean_ctor_get(x_2, 1); +lean_inc(x_453); +lean_dec(x_2); +x_2 = x_453; +x_3 = x_340; +x_8 = x_452; +goto _start; +} +case 11: +{ +lean_object* x_455; lean_object* x_456; +lean_dec(x_341); +x_455 = lean_ctor_get(x_342, 1); +lean_inc(x_455); +lean_dec(x_342); +x_456 = lean_ctor_get(x_2, 2); +lean_inc(x_456); +lean_dec(x_2); +x_2 = x_456; +x_3 = x_340; +x_8 = x_455; +goto _start; +} +default: +{ +lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; +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_458 = lean_ctor_get(x_342, 1); +lean_inc(x_458); +if (lean_is_exclusive(x_342)) { + lean_ctor_release(x_342, 0); + lean_ctor_release(x_342, 1); + x_459 = x_342; +} else { + lean_dec_ref(x_342); + x_459 = lean_box(0); +} +x_460 = lean_box(0); +if (lean_is_scalar(x_341)) { + x_461 = lean_alloc_ctor(0, 2, 0); +} else { + x_461 = x_341; +} +lean_ctor_set(x_461, 0, x_460); +lean_ctor_set(x_461, 1, x_340); +if (lean_is_scalar(x_459)) { + x_462 = lean_alloc_ctor(0, 2, 0); +} else { + x_462 = x_459; +} +lean_ctor_set(x_462, 0, x_461); +lean_ctor_set(x_462, 1, x_458); +return x_462; +} +} +} +else +{ +lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_463 = lean_ctor_get(x_342, 1); +lean_inc(x_463); +if (lean_is_exclusive(x_342)) { + lean_ctor_release(x_342, 0); + lean_ctor_release(x_342, 1); + x_464 = x_342; +} else { + lean_dec_ref(x_342); + x_464 = lean_box(0); +} +x_465 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_465, 0, x_2); +if (lean_is_scalar(x_341)) { + x_466 = lean_alloc_ctor(0, 2, 0); +} else { + x_466 = x_341; +} +lean_ctor_set(x_466, 0, x_465); +lean_ctor_set(x_466, 1, x_340); +if (lean_is_scalar(x_464)) { + x_467 = lean_alloc_ctor(0, 2, 0); +} else { + x_467 = x_464; +} +lean_ctor_set(x_467, 0, x_466); +lean_ctor_set(x_467, 1, x_463); +return x_467; +} +} +else +{ +lean_object* x_468; lean_object* x_469; lean_object* x_470; lean_object* x_471; +lean_dec(x_341); +lean_dec(x_340); +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_468 = lean_ctor_get(x_342, 0); +lean_inc(x_468); +x_469 = lean_ctor_get(x_342, 1); +lean_inc(x_469); +if (lean_is_exclusive(x_342)) { + lean_ctor_release(x_342, 0); + lean_ctor_release(x_342, 1); + x_470 = x_342; +} else { + lean_dec_ref(x_342); + x_470 = lean_box(0); +} +if (lean_is_scalar(x_470)) { + x_471 = lean_alloc_ctor(1, 2, 0); +} else { + x_471 = x_470; +} +lean_ctor_set(x_471, 0, x_468); +lean_ctor_set(x_471, 1, x_469); +return x_471; +} +} +} +} +} +} +static lean_object* _init_l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(64u); +x_2 = l_Lean_mkHashSetImp___rarg(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___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; lean_object* x_9; +x_8 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___spec__1___closed__1; +x_9 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2(x_1, x_2, x_8, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = lean_ctor_get(x_13, 0); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +return x_9; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_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_Meta_occursOrInType___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_inc(x_1); +x_9 = lean_infer_type(x_1, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_expr_eqv(x_1, x_2); +lean_dec(x_1); +if (x_12 == 0) +{ +uint8_t x_13; lean_object* x_14; +x_13 = l_Lean_Expr_occurs(x_2, x_11); +x_14 = lean_box(x_13); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +uint8_t x_15; lean_object* x_16; +lean_dec(x_11); +lean_dec(x_2); +x_15 = 1; +x_16 = lean_box(x_15); +lean_ctor_set(x_9, 0, x_16); +return x_9; +} +} +else +{ +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_9, 0); +x_18 = lean_ctor_get(x_9, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_9); +x_19 = lean_expr_eqv(x_1, x_2); +lean_dec(x_1); +if (x_19 == 0) +{ +uint8_t x_20; lean_object* x_21; lean_object* x_22; +x_20 = l_Lean_Expr_occurs(x_2, 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; +} +else +{ +uint8_t x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_17); +lean_dec(x_2); +x_23 = 1; +x_24 = lean_box(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_18); +return x_25; +} +} +} +else +{ +uint8_t x_26; +lean_dec(x_2); +lean_dec(x_1); +x_26 = !lean_is_exclusive(x_9); +if (x_26 == 0) +{ +return x_9; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_9, 0); +x_28 = lean_ctor_get(x_9, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_9); +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_EXPORT lean_object* l_Lean_Meta_occursOrInType___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +uint8_t x_8; +x_8 = l_Lean_Expr_isFVar(x_2); +if (x_8 == 0) +{ +uint8_t x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_9 = lean_expr_eqv(x_2, x_1); +lean_dec(x_1); +lean_dec(x_2); +x_10 = lean_box(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_7); +return x_11; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = l_Lean_Meta_occursOrInType___lambda__1(x_2, x_1, x_12, x_3, x_4, x_5, x_6, x_7); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_occursOrInType(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_alloc_closure((void*)(l_Lean_Meta_occursOrInType___lambda__2), 7, 1); +lean_closure_set(x_8, 0, x_1); +x_9 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Meta_occursOrInType___spec__1(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_9); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_ctor_get(x_9, 0); +lean_dec(x_12); +x_13 = 0; +x_14 = lean_box(x_13); +lean_ctor_set(x_9, 0, x_14); +return x_9; +} +else +{ +lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_dec(x_9); +x_16 = 0; +x_17 = lean_box(x_16); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; +} +} +else +{ +uint8_t x_19; +lean_dec(x_10); +x_19 = !lean_is_exclusive(x_9); +if (x_19 == 0) +{ +lean_object* x_20; uint8_t x_21; lean_object* x_22; +x_20 = lean_ctor_get(x_9, 0); +lean_dec(x_20); +x_21 = 1; +x_22 = lean_box(x_21); +lean_ctor_set(x_9, 0, x_22); +return x_9; +} +else +{ +lean_object* x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_9, 1); +lean_inc(x_23); +lean_dec(x_9); +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; +} +} +} +else +{ +uint8_t x_27; +x_27 = !lean_is_exclusive(x_9); +if (x_27 == 0) +{ +return x_9; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_9, 0); +x_29 = lean_ctor_get(x_9, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_9); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Expr_FindImpl_findM_x3f_visit___at_Lean_Meta_occursOrInType___spec__2___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_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_occursOrInType___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Meta_occursOrInType___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_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___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: { @@ -624,7 +3705,7 @@ x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); if (lean_obj_tag(x_19) == 7) { -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_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_18, 1); lean_inc(x_20); lean_dec(x_18); @@ -636,98 +3717,153 @@ x_23 = lean_ctor_get(x_19, 2); lean_inc(x_23); lean_dec(x_19); x_24 = lean_array_fget(x_3, x_6); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_4); lean_inc(x_24); -x_25 = l_Lean_Expr_occurs(x_24, x_4); -if (x_25 == 0) +x_25 = l_Lean_Meta_occursOrInType(x_24, x_4, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_unbox(x_26); +lean_dec(x_26); +if (x_27 == 0) { -lean_object* x_26; lean_object* x_27; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_dec(x_24); -x_26 = lean_box(x_2); -x_27 = lean_alloc_closure((void*)(l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___lambda__1___boxed), 15, 9); -lean_closure_set(x_27, 0, x_6); -lean_closure_set(x_27, 1, x_23); -lean_closure_set(x_27, 2, x_8); -lean_closure_set(x_27, 3, x_9); -lean_closure_set(x_27, 4, x_1); -lean_closure_set(x_27, 5, x_26); -lean_closure_set(x_27, 6, x_3); -lean_closure_set(x_27, 7, x_4); -lean_closure_set(x_27, 8, x_5); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_box(x_2); +x_30 = lean_alloc_closure((void*)(l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___lambda__1___boxed), 15, 9); +lean_closure_set(x_30, 0, x_6); +lean_closure_set(x_30, 1, x_23); +lean_closure_set(x_30, 2, x_8); +lean_closure_set(x_30, 3, x_9); +lean_closure_set(x_30, 4, x_1); +lean_closure_set(x_30, 5, x_29); +lean_closure_set(x_30, 6, x_3); +lean_closure_set(x_30, 7, x_4); +lean_closure_set(x_30, 8, x_5); if (x_2 == 0) { -uint8_t x_28; uint8_t x_29; lean_object* x_30; -x_28 = 1; -x_29 = 0; -x_30 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_21, x_28, x_22, x_27, x_29, x_10, x_11, x_12, x_13, x_20); -return x_30; -} -else -{ uint8_t x_31; uint8_t x_32; lean_object* x_33; -x_31 = 0; +x_31 = 1; x_32 = 0; -x_33 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_21, x_31, x_22, x_27, x_32, x_10, x_11, x_12, x_13, x_20); +x_33 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_21, x_31, x_22, x_30, x_32, x_10, x_11, x_12, x_13, x_28); return x_33; } +else +{ +uint8_t x_34; uint8_t x_35; lean_object* x_36; +x_34 = 0; +x_35 = 0; +x_36 = l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2___rarg(x_21, x_34, x_22, x_30, x_35, x_10, x_11, x_12, x_13, x_28); +return x_36; +} } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_dec(x_22); lean_dec(x_21); -x_34 = lean_unsigned_to_nat(1u); -x_35 = lean_nat_add(x_6, x_34); +x_37 = lean_ctor_get(x_25, 1); +lean_inc(x_37); +lean_dec(x_25); +x_38 = lean_unsigned_to_nat(1u); +x_39 = lean_nat_add(x_6, x_38); lean_dec(x_6); -x_36 = lean_expr_instantiate1(x_23, x_24); +x_40 = lean_expr_instantiate1(x_23, x_24); lean_dec(x_23); -x_37 = lean_array_push(x_8, x_24); -x_6 = x_35; -x_7 = x_36; -x_8 = x_37; -x_14 = x_20; +x_41 = lean_array_push(x_8, x_24); +x_6 = x_39; +x_7 = x_40; +x_8 = x_41; +x_14 = x_37; goto _start; } } 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_dec(x_19); +uint8_t x_43; +lean_dec(x_24); +lean_dec(x_23); +lean_dec(x_22); +lean_dec(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_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_39 = lean_ctor_get(x_18, 1); -lean_inc(x_39); -lean_dec(x_18); -x_40 = lean_ctor_get(x_1, 0); -lean_inc(x_40); lean_dec(x_1); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_MessageData_ofName(x_41); -x_43 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__2; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -x_45 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__4; -x_46 = lean_alloc_ctor(7, 2, 0); +x_43 = !lean_is_exclusive(x_25); +if (x_43 == 0) +{ +return x_25; +} +else +{ +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_25, 0); +x_45 = lean_ctor_get(x_25, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_25); +x_46 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_46, 0, x_44); lean_ctor_set(x_46, 1, x_45); -x_47 = l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___spec__1(x_46, x_10, x_11, x_12, x_13, x_39); +return x_46; +} +} +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_dec(x_19); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_47 = lean_ctor_get(x_18, 1); +lean_inc(x_47); +lean_dec(x_18); +x_48 = lean_ctor_get(x_1, 0); +lean_inc(x_48); +lean_dec(x_1); +x_49 = lean_ctor_get(x_48, 0); +lean_inc(x_49); +lean_dec(x_48); +x_50 = l_Lean_MessageData_ofName(x_49); +x_51 = l___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__2; +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___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___closed__4; +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 = l_Lean_throwError___at___private_Lean_Meta_Injective_0__Lean_Meta_mkInjectiveTheoremTypeCore_x3f_mkArgs2___spec__1(x_54, x_10, x_11, x_12, x_13, x_47); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); -return x_47; +return x_55; } } else { -uint8_t x_48; +uint8_t x_56; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -739,23 +3875,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_18); -if (x_48 == 0) +x_56 = !lean_is_exclusive(x_18); +if (x_56 == 0) { return x_18; } else { -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_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_18, 0); +x_58 = lean_ctor_get(x_18, 1); +lean_inc(x_58); +lean_inc(x_57); 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; +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; } } } @@ -2118,7 +5254,7 @@ lean_closure_set(x_21, 0, x_8); lean_closure_set(x_21, 1, x_11); lean_closure_set(x_21, 2, x_20); lean_closure_set(x_21, 3, x_1); -x_22 = l_Lean_Meta_forallBoundedTelescope___at_Lean_Meta_addPPExplicitToExposeDiff_visit___spec__2___rarg(x_16, x_19, x_21, x_3, x_4, x_5, x_6, x_17); +x_22 = l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___rarg(x_16, x_19, x_21, x_3, x_4, x_5, x_6, x_17); return x_22; } else @@ -2558,7 +5694,7 @@ static lean_object* _init_l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOf lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__1; x_2 = l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__2; -x_3 = lean_unsigned_to_nat(87u); +x_3 = lean_unsigned_to_nat(116u); x_4 = lean_unsigned_to_nat(30u); x_5 = l___private_Lean_Meta_Injective_0__Lean_Meta_solveEqOfCtorEq___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4141,7 +7277,7 @@ return x_87; } } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__1() { _start: { lean_object* x_1; @@ -4149,17 +7285,17 @@ x_1 = lean_mk_string_from_bytes("genInjectivity", 14); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____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_Injective___hyg_1818____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__3() { _start: { lean_object* x_1; @@ -4167,13 +7303,13 @@ x_1 = lean_mk_string_from_bytes("generate injectivity theorems for inductive dat return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__4() { _start: { uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = 1; x_2 = l___private_Lean_Meta_Injective_0__Lean_Meta_throwInjectiveTheoremFailure___rarg___closed__1; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__3; x_4 = lean_box(x_1); x_5 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_5, 0, x_4); @@ -4182,7 +7318,7 @@ lean_ctor_set(x_5, 2, x_3); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__5() { _start: { lean_object* x_1; @@ -4190,7 +7326,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__6() { _start: { lean_object* x_1; @@ -4198,25 +7334,25 @@ x_1 = lean_mk_string_from_bytes("Meta", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____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_Injective___hyg_1818____closed__5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__6; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__6; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__4; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1818____closed__7; -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); +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__4; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Injective___hyg_1958____closed__7; +x_5 = l_Lean_Option_register___at_Lean_Elab_initFn____x40_Lean_Elab_AutoBound___hyg_7____spec__1(x_2, x_3, x_4, x_1); return x_5; } } @@ -4633,180 +7769,68 @@ lean_object* x_16; lean_object* x_17; x_16 = lean_ctor_get(x_13, 1); lean_inc(x_16); lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); -x_20 = lean_ctor_get(x_8, 2); -x_21 = lean_ctor_get(x_8, 4); -x_22 = lean_ctor_get(x_8, 5); -x_23 = lean_ctor_get(x_8, 6); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_8); -x_24 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___closed__3; -x_25 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_25, 0, x_18); -lean_ctor_set(x_25, 1, x_19); -lean_ctor_set(x_25, 2, x_20); -lean_ctor_set(x_25, 3, x_24); -lean_ctor_set(x_25, 4, x_21); -lean_ctor_set(x_25, 5, x_22); -lean_ctor_set(x_25, 6, x_23); -x_26 = lean_st_ref_set(x_1, x_25, x_9); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_28 = x_26; -} else { - lean_dec_ref(x_26); - x_28 = lean_box(0); -} -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_28; -} -lean_ctor_set(x_29, 0, x_6); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___boxed), 2, 0); -return x_4; -} -} -static double _init_l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___spec__8___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___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_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; double x_18; double x_19; double x_20; double x_21; double x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_15 = lean_ctor_get(x_13, 0); -x_16 = 0; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Float_ofScientific(x_8, x_16, x_17); -lean_dec(x_8); -x_19 = l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___spec__8___closed__1; -x_20 = lean_float_div(x_18, x_19); -x_21 = l_Float_ofScientific(x_15, x_16, x_17); -lean_dec(x_15); -x_22 = lean_float_div(x_21, x_19); -x_23 = lean_box_float(x_20); -x_24 = lean_box_float(x_22); -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, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_13, 0, x_26); -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_dec(x_13); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_8, x_29, x_30); -lean_dec(x_8); -x_32 = l_Lean_withStartStopSeconds___at_Lean_Meta_mkInjectiveTheorems___spec__8___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = l_Float_ofScientific(x_27, x_29, x_30); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_32); -x_36 = lean_box_float(x_33); -x_37 = lean_box_float(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_11); -lean_ctor_set(x_39, 1, 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_28); -return x_40; +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } else { -uint8_t x_41; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); lean_dec(x_8); -x_41 = !lean_is_exclusive(x_10); -if (x_41 == 0) -{ -return x_10; +x_24 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___closed__3; +x_25 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_19); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; +} +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); +return x_29; } -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_10, 0); -x_43 = lean_ctor_get(x_10, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_10); -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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_mkInjectiveTheorems___spec__7___rarg___boxed), 2, 0); +return x_4; } } -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, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjectiveTheorems___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) { _start: { uint8_t x_10; @@ -5075,7 +8099,7 @@ return x_99; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___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) { _start: { if (lean_obj_tag(x_1) == 0) @@ -5100,89 +8124,23 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___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) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_6); lean_inc(x_9); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__9(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__8(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__10(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_Meta_mkInjectiveTheorems___spec__9(x_4, x_7, x_8, x_9, x_10, x_13); lean_dec(x_9); lean_dec(x_4); return x_14; } } -static double _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -5192,7 +8150,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -5200,11 +8158,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___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, uint8_t x_8, double x_9, double 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: { double x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1; +x_17 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_18 = lean_alloc_ctor(0, 2, 17); @@ -5213,7 +8171,7 @@ lean_ctor_set(x_18, 1, x_3); lean_ctor_set_float(x_18, sizeof(void*)*2, x_17); lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_17); lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_2); -x_19 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__2; +x_19 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2___closed__2; x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_19); if (x_20 == 0) { @@ -5223,7 +8181,7 @@ lean_object* x_21; lean_object* x_22; lean_dec(x_3); lean_dec(x_1); x_21 = lean_box(0); -x_22 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); return x_22; } else @@ -5237,7 +8195,7 @@ lean_ctor_set_float(x_23, sizeof(void*)*2, x_9); lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_2); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); +x_25 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); return x_25; } } @@ -5252,12 +8210,12 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); return x_28; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_mkInjectiveTheorems___spec__5___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -5265,16 +8223,16 @@ x_1 = lean_mk_string_from_bytes(" #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -621,6 +621,7 @@ lean_object* initialize_Lean_Meta_Tactic_Simp_Simproc(uint8_t builtin, lean_obje lean_object* initialize_Lean_Meta_Tactic_Simp_BuiltinSimprocs(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_RegisterCommand(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_Attr(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Simp_Diagnostics(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Simp(uint8_t builtin, lean_object* w) { lean_object * res; @@ -656,6 +657,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_Attr(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Diagnostics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); l_Lean_initFn____x40_Lean_Meta_Tactic_Simp___hyg_4____closed__1 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Simp___hyg_4____closed__1(); lean_mark_persistent(l_Lean_initFn____x40_Lean_Meta_Tactic_Simp___hyg_4____closed__1); l_Lean_initFn____x40_Lean_Meta_Tactic_Simp___hyg_4____closed__2 = _init_l_Lean_initFn____x40_Lean_Meta_Tactic_Simp___hyg_4____closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c index b9350c58109e..c28d5831f58c 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c @@ -2937,7 +2937,7 @@ x_8 = l_Lean_Meta_getSimpCongrTheorems___rarg(x_2, x_7); x_9 = !lean_is_exclusive(x_8); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint32_t x_14; lean_object* x_15; uint32_t x_16; lean_object* x_17; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint32_t x_14; lean_object* x_15; uint32_t x_16; lean_object* x_17; lean_object* x_18; x_10 = lean_ctor_get(x_8, 0); x_11 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__3; x_12 = lean_array_push(x_11, x_6); @@ -2945,41 +2945,45 @@ x_13 = lean_box(0); x_14 = 0; x_15 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__1; x_16 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__2; -x_17 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_12); -lean_ctor_set(x_17, 2, x_10); -lean_ctor_set(x_17, 3, x_13); -lean_ctor_set_uint32(x_17, sizeof(void*)*4, x_16); -lean_ctor_set_uint32(x_17, sizeof(void*)*4 + 4, x_14); -lean_ctor_set(x_8, 0, x_17); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_18, 0, x_15); +lean_ctor_set(x_18, 1, x_12); +lean_ctor_set(x_18, 2, x_10); +lean_ctor_set(x_18, 3, x_13); +lean_ctor_set(x_18, 4, x_17); +lean_ctor_set_uint32(x_18, sizeof(void*)*5, x_16); +lean_ctor_set_uint32(x_18, sizeof(void*)*5 + 4, x_14); +lean_ctor_set(x_8, 0, x_18); return x_8; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint32_t x_23; lean_object* x_24; uint32_t x_25; lean_object* x_26; lean_object* x_27; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint32_t x_24; lean_object* x_25; uint32_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_inc(x_18); lean_dec(x_8); -x_20 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__3; -x_21 = lean_array_push(x_20, x_6); -x_22 = lean_box(0); -x_23 = 0; -x_24 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__1; -x_25 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__2; -x_26 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_21); -lean_ctor_set(x_26, 2, x_18); -lean_ctor_set(x_26, 3, x_22); -lean_ctor_set_uint32(x_26, sizeof(void*)*4, x_25); -lean_ctor_set_uint32(x_26, sizeof(void*)*4 + 4, x_23); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_19); -return x_27; +x_21 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__3; +x_22 = lean_array_push(x_21, x_6); +x_23 = lean_box(0); +x_24 = 0; +x_25 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__1; +x_26 = l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__2; +x_27 = lean_unsigned_to_nat(0u); +x_28 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_28, 0, x_25); +lean_ctor_set(x_28, 1, x_22); +lean_ctor_set(x_28, 2, x_19); +lean_ctor_set(x_28, 3, x_23); +lean_ctor_set(x_28, 4, x_27); +lean_ctor_set_uint32(x_28, sizeof(void*)*5, x_26); +lean_ctor_set_uint32(x_28, sizeof(void*)*5 + 4, x_24); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_28); +lean_ctor_set(x_29, 1, x_20); +return x_29; } } } diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Nat.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Nat.c index 95595895c2d3..a0250bbc86a4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Nat.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/BuiltinSimprocs/Nat.c @@ -12572,9 +12572,9 @@ lean_dec(x_42); x_48 = l_Lean_mkNatLit(x_47); lean_inc(x_43); x_49 = l___private_Lean_Meta_Tactic_Simp_BuiltinSimprocs_Nat_0__Nat_mkLENat(x_48, x_43); -lean_inc(x_44); lean_inc(x_13); -x_50 = l___private_Lean_Meta_Tactic_Simp_BuiltinSimprocs_Nat_0__Nat_mkLENat(x_13, x_44); +lean_inc(x_44); +x_50 = l___private_Lean_Meta_Tactic_Simp_BuiltinSimprocs_Nat_0__Nat_mkLENat(x_44, x_13); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c new file mode 100644 index 000000000000..e202895289e2 --- /dev/null +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Diagnostics.c @@ -0,0 +1,1955 @@ +// Lean compiler output +// Module: Lean.Meta.Tactic.Simp.Diagnostics +// Imports: Lean.Meta.Diagnostics Lean.Meta.Tactic.Simp.Types +#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 +static lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1___boxed(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1; +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__5; +lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__2; +lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Simp_recordTriedSimpTheorem___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8(lean_object*, lean_object*, size_t, size_t, 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_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t lean_usize_dec_eq(size_t, size_t); +lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5; +lean_object* lean_array_fget(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Meta_DiagSummary_isEmpty(lean_object*); +lean_object* l_Lean_Meta_instDecidableLtOrigin___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__7; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11; +extern lean_object* l_instInhabitedNat; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7; +lean_object* l_Lean_Meta_mkDiagSummary(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_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Origin_key(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1; +lean_object* lean_st_ref_get(lean_object*, lean_object*); +extern lean_object* l_Lean_MessageData_nil; +LEAN_EXPORT uint8_t l_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3; +lean_object* l_Lean_MessageData_ofConst(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___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_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___closed__5; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_appendSection(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___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_mkConstWithLevelParams___at_Lean_Meta_mkSimpCongrTheorem___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofExpr(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_crossEmoji; +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__6; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__3(lean_object*, lean_object*, lean_object*); +uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10; +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__3; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___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_Meta_Simp_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__1; +lean_object* lean_nat_sub(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__8; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_diagnostics_threshold; +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2(lean_object*, lean_object*, lean_object*); +size_t lean_usize_add(size_t, size_t); +extern lean_object* l_Lean_Meta_instInhabitedOrigin; +lean_object* lean_array_uget(lean_object*, size_t); +lean_object* l_Lean_Expr_fvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7___lambda__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); +lean_object* lean_array_get_size(lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4; +uint8_t lean_usize_dec_lt(size_t, size_t); +lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3; +static lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2; +lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_ofName(lean_object*); +lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Lean_Meta_Simp_reportDiag___closed__4; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Array_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___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; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) +{ +lean_object* x_7; +x_7 = lean_array_uget(x_2, x_3); +switch (lean_obj_tag(x_7)) { +case 0: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +lean_inc(x_1); +x_10 = lean_apply_3(x_1, x_5, x_8, x_9); +if (lean_obj_tag(x_10) == 0) +{ +uint8_t x_11; +lean_dec(x_1); +x_11 = !lean_is_exclusive(x_10); +if (x_11 == 0) +{ +return x_10; +} +else +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_10, 0); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +return x_13; +} +} +else +{ +lean_object* x_14; size_t x_15; size_t x_16; +x_14 = lean_ctor_get(x_10, 0); +lean_inc(x_14); +lean_dec(x_10); +x_15 = 1; +x_16 = lean_usize_add(x_3, x_15); +x_3 = x_16; +x_5 = x_14; +goto _start; +} +} +case 1: +{ +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_7, 0); +lean_inc(x_18); +lean_dec(x_7); +lean_inc(x_1); +x_19 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4___rarg(x_1, x_18, x_5); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +lean_dec(x_1); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +return x_19; +} +else +{ +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_22, 0, x_21); +return x_22; +} +} +else +{ +lean_object* x_23; size_t x_24; size_t x_25; +x_23 = lean_ctor_get(x_19, 0); +lean_inc(x_23); +lean_dec(x_19); +x_24 = 1; +x_25 = lean_usize_add(x_3, x_24); +x_3 = x_25; +x_5 = x_23; +goto _start; +} +} +default: +{ +size_t x_27; size_t x_28; +x_27 = 1; +x_28 = lean_usize_add(x_3, x_27); +x_3 = x_28; +goto _start; +} +} +} +else +{ +lean_object* x_30; +lean_dec(x_1); +x_30 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_30, 0, x_5); +return x_30; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___rarg___boxed), 5, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_lt(x_5, x_7); +lean_dec(x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_dec(x_5); +lean_dec(x_1); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_6); +return x_9; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_array_fget(x_2, x_5); +x_11 = lean_array_fget(x_3, x_5); +lean_inc(x_1); +x_12 = lean_apply_3(x_1, x_6, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +uint8_t x_13; +lean_dec(x_5); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +return x_12; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_inc(x_14); +lean_dec(x_12); +x_15 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_15, 0, x_14); +return x_15; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_add(x_5, x_17); +lean_dec(x_5); +x_4 = lean_box(0); +x_5 = x_18; +x_6 = x_16; +goto _start; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg___boxed), 6, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_4 = lean_ctor_get(x_2, 0); +lean_inc(x_4); +lean_dec(x_2); +x_5 = lean_array_get_size(x_4); +x_6 = lean_unsigned_to_nat(0u); +x_7 = lean_nat_dec_lt(x_6, x_5); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_3); +return x_8; +} +else +{ +uint8_t x_9; +x_9 = lean_nat_dec_le(x_5, x_5); +if (x_9 == 0) +{ +lean_object* x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_3); +return x_10; +} +else +{ +size_t x_11; size_t x_12; lean_object* x_13; +x_11 = 0; +x_12 = lean_usize_of_nat(x_5); +lean_dec(x_5); +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___rarg(x_1, x_4, x_11, x_12, x_3); +lean_dec(x_4); +return x_13; +} +} +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_2, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_2, 1); +lean_inc(x_15); +lean_dec(x_2); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg(x_1, x_14, x_15, lean_box(0), x_16, x_3); +lean_dec(x_15); +lean_dec(x_14); +return x_17; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = lean_alloc_closure((void*)(l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4___rarg), 3, 0); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; lean_object* x_5; +x_4 = lean_ctor_get(x_1, 0); +lean_inc(x_4); +lean_dec(x_1); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__4___rarg(x_2, x_4, x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_5, 0, x_3); +lean_ctor_set(x_5, 1, x_4); +x_6 = lean_apply_2(x_1, x_5, x_2); +if (lean_obj_tag(x_6) == 0) +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_8, 0, x_7); +return x_8; +} +else +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_6, 0); +lean_inc(x_9); +lean_dec(x_6); +x_10 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_10, 0, x_9); +return x_10; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2(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_alloc_closure((void*)(l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2___lambda__1), 4, 1); +lean_closure_set(x_4, 0, x_3); +x_5 = l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__3(x_1, x_4, x_2); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7___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; lean_object* x_7; uint8_t x_8; +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_6 = lean_ctor_get(x_3, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_3, 1); +lean_inc(x_7); +lean_dec(x_3); +x_8 = lean_nat_dec_eq(x_5, x_7); +if (x_8 == 0) +{ +uint8_t x_9; lean_object* x_10; +lean_dec(x_6); +lean_dec(x_4); +lean_dec(x_1); +x_9 = lean_nat_dec_lt(x_7, x_5); +lean_dec(x_5); +lean_dec(x_7); +x_10 = lean_box(x_9); +return x_10; +} +else +{ +lean_object* x_11; +lean_dec(x_7); +lean_dec(x_5); +x_11 = lean_apply_2(x_1, x_4, x_6); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7(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; +lean_inc(x_1); +x_5 = lean_alloc_closure((void*)(l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7___lambda__1), 3, 1); +lean_closure_set(x_5, 0, x_1); +x_6 = lean_nat_dec_lt(x_3, x_4); +if (x_6 == 0) +{ +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +return x_2; +} +else +{ +lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_inc(x_3); +x_7 = l_Array_qpartition___rarg(x_2, x_5, x_3, x_4); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_nat_dec_le(x_4, x_8); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_inc(x_1); +x_11 = l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7(x_1, x_9, x_3, x_8); +x_12 = lean_unsigned_to_nat(1u); +x_13 = lean_nat_add(x_8, x_12); +lean_dec(x_8); +x_2 = x_11; +x_3 = x_13; +goto _start; +} +else +{ +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_1); +return x_9; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_3); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_6 = lean_ctor_get(x_3, 0); +x_7 = lean_ctor_get(x_3, 1); +x_8 = lean_nat_dec_lt(x_1, x_7); +if (x_8 == 0) +{ +lean_object* x_9; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +x_9 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_9, 0, x_4); +return x_9; +} +else +{ +lean_object* x_10; uint8_t x_11; +lean_inc(x_6); +x_10 = lean_apply_1(x_2, x_6); +x_11 = lean_unbox(x_10); +lean_dec(x_10); +if (x_11 == 0) +{ +lean_object* x_12; +lean_free_object(x_3); +lean_dec(x_7); +lean_dec(x_6); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_4); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_array_push(x_4, x_3); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_13); +return x_14; +} +} +} +else +{ +lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_15 = lean_ctor_get(x_3, 0); +x_16 = lean_ctor_get(x_3, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_3); +x_17 = lean_nat_dec_lt(x_1, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_2); +x_18 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_18, 0, x_4); +return x_18; +} +else +{ +lean_object* x_19; uint8_t x_20; +lean_inc(x_15); +x_19 = lean_apply_1(x_2, x_15); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_16); +lean_dec(x_15); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_4); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_22, 0, x_15); +lean_ctor_set(x_22, 1, x_16); +x_23 = lean_array_push(x_4, x_22); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_23); +return x_24; +} +} +} +} +} +static lean_object* _init_l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(0u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_5 = lean_alloc_closure((void*)(l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1___boxed), 4, 2); +lean_closure_set(x_5, 0, x_2); +lean_closure_set(x_5, 1, x_3); +x_6 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1; +x_7 = l_Lean_PersistentHashMap_forIn___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__2(x_1, x_6, x_5); +x_8 = lean_array_get_size(x_7); +x_9 = lean_unsigned_to_nat(1u); +x_10 = lean_nat_sub(x_8, x_9); +lean_dec(x_8); +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7(x_4, x_7, x_11, x_10); +lean_dec(x_10); +return x_12; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ↦ ", 5); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("\n", 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6; +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_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ", 2); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10; +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_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_3 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +uint8_t x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = l_Array_isEmpty___rarg(x_3); +x_11 = l___private_Init_Data_Repr_0__Nat_reprFast(x_1); +x_12 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_12, 0, x_11); +x_13 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = l_Lean_stringToMessageData(x_4); +if (x_10 == 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; +x_15 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_2); +x_17 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_13); +x_20 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_21 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_21, 0, x_19); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_14); +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +x_24 = lean_array_push(x_3, x_23); +x_25 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_25, 0, x_24); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_9); +return x_26; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_27 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12; +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_2); +x_29 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4; +x_30 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_13); +x_32 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_14); +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_32); +x_36 = lean_array_push(x_3, x_35); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_9); +return x_38; +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" ", 1); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1; +x_2 = l_Lean_crossEmoji; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1; +x_3 = lean_string_append(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(", succeeded: ", 13); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +if (lean_obj_tag(x_2) == 0) +{ +lean_object* x_11; lean_object* x_12; +lean_dec(x_3); +x_11 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1; +x_12 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1(x_1, x_5, 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; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = lean_ctor_get(x_2, 0); +lean_inc(x_13); +lean_dec(x_2); +x_14 = l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Simp_recordTriedSimpTheorem___spec__1(x_13, x_3); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; +x_15 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3; +x_16 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1(x_1, x_5, x_4, x_15, 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_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_17 = lean_ctor_get(x_14, 0); +lean_inc(x_17); +lean_dec(x_14); +x_18 = l___private_Init_Data_Repr_0__Nat_reprFast(x_17); +x_19 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4; +x_20 = lean_string_append(x_19, x_18); +lean_dec(x_18); +x_21 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1; +x_22 = lean_string_append(x_20, x_21); +x_23 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1(x_1, x_5, x_4, x_22, 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_22); +return x_23; +} +} +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes(" (builtin simproc)", 18); +return x_1; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = lean_usize_dec_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_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_13; lean_object* x_14; lean_object* x_15; lean_object* x_21; +x_13 = lean_array_uget(x_2, x_4); +x_21 = lean_ctor_get(x_13, 0); +lean_inc(x_21); +switch (lean_obj_tag(x_21)) { +case 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; uint8_t x_28; +x_22 = lean_ctor_get(x_13, 1); +lean_inc(x_22); +lean_dec(x_13); +x_23 = lean_ctor_get(x_21, 0); +lean_inc(x_23); +x_24 = lean_st_ref_get(x_9, x_10); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = lean_ctor_get(x_25, 0); +lean_inc(x_27); +lean_dec(x_25); +lean_inc(x_23); +x_28 = l_Lean_Environment_contains(x_27, x_23); +if (x_28 == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_29 = l_Lean_MessageData_ofName(x_23); +x_30 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_31 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2; +x_33 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_34 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2(x_22, x_1, x_21, x_5, x_33, x_6, x_7, x_8, x_9, x_26); +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_14 = x_35; +x_15 = x_36; +goto block_20; +} +else +{ +lean_object* x_37; +x_37 = l_Lean_mkConstWithLevelParams___at_Lean_Meta_mkSimpCongrTheorem___spec__1(x_23, x_6, x_7, x_8, x_9, x_26); +if (lean_obj_tag(x_37) == 0) +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +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_MessageData_ofConst(x_38); +x_41 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_40); +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_44 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2(x_22, x_1, x_21, x_5, x_43, x_6, x_7, x_8, x_9, x_39); +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_14 = x_45; +x_15 = x_46; +goto block_20; +} +else +{ +uint8_t x_47; +lean_dec(x_22); +lean_dec(x_21); +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_47 = !lean_is_exclusive(x_37); +if (x_47 == 0) +{ +return x_37; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_37, 0); +x_49 = lean_ctor_get(x_37, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_37); +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; +} +} +} +} +case 1: +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_51 = lean_ctor_get(x_13, 1); +lean_inc(x_51); +lean_dec(x_13); +x_52 = lean_ctor_get(x_21, 0); +lean_inc(x_52); +x_53 = l_Lean_Expr_fvar___override(x_52); +x_54 = l_Lean_MessageData_ofExpr(x_53); +x_55 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2; +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +x_57 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_58 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2(x_51, x_1, x_21, x_5, x_57, x_6, x_7, x_8, x_9, x_10); +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_14 = x_59; +x_15 = x_60; +goto block_20; +} +default: +{ +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_61 = lean_ctor_get(x_13, 1); +lean_inc(x_61); +lean_dec(x_13); +x_62 = l_Lean_Meta_Origin_key(x_21); +x_63 = l_Lean_MessageData_ofName(x_62); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_1); +x_64 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2(x_61, x_1, x_21, x_5, x_63, x_6, x_7, x_8, x_9, x_10); +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_14 = x_65; +x_15 = x_66; +goto block_20; +} +} +block_20: +{ +lean_object* x_16; size_t x_17; size_t x_18; +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = 1; +x_18 = lean_usize_add(x_4, x_17); +x_4 = x_18; +x_5 = x_16; +x_10 = x_15; +goto _start; +} +} +} +} +LEAN_EXPORT uint8_t l_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1(lean_object* x_1) { +_start: +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics_threshold; +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_instDecidableLtOrigin___boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_instInhabitedOrigin; +x_2 = l_instInhabitedNat; +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1; +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_EXPORT lean_object* l_Lean_Meta_Simp_mkSimpDiagSummary(lean_object* x_1, 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; uint8_t x_14; +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +x_9 = l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1; +x_10 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_8, x_9); +lean_dec(x_8); +x_11 = l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2; +x_12 = l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3; +x_13 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1(x_1, x_10, x_11, x_12); +x_14 = l_Array_isEmpty___rarg(x_13); +if (x_14 == 0) +{ +lean_object* x_15; size_t x_16; size_t x_17; lean_object* x_18; lean_object* x_19; +x_15 = lean_array_get_size(x_13); +x_16 = lean_usize_of_nat(x_15); +x_17 = 0; +x_18 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1; +x_19 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8(x_2, x_13, x_16, x_17, x_18, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_19) == 0) +{ +uint8_t x_20; +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_21 = lean_ctor_get(x_19, 0); +x_22 = lean_unsigned_to_nat(0u); +x_23 = lean_nat_dec_lt(x_22, x_15); +lean_dec(x_15); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_13); +x_24 = l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4; +x_25 = l___private_Init_GetElem_0__outOfBounds___rarg(x_24); +x_26 = lean_ctor_get(x_25, 1); +lean_inc(x_26); +lean_dec(x_25); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_21); +lean_ctor_set(x_27, 1, x_26); +lean_ctor_set(x_19, 0, x_27); +return x_19; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_array_fget(x_13, x_22); +lean_dec(x_13); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_21); +lean_ctor_set(x_30, 1, x_29); +lean_ctor_set(x_19, 0, x_30); +return x_19; +} +} +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_31 = lean_ctor_get(x_19, 0); +x_32 = lean_ctor_get(x_19, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_19); +x_33 = lean_unsigned_to_nat(0u); +x_34 = lean_nat_dec_lt(x_33, x_15); +lean_dec(x_15); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_13); +x_35 = l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4; +x_36 = l___private_Init_GetElem_0__outOfBounds___rarg(x_35); +x_37 = lean_ctor_get(x_36, 1); +lean_inc(x_37); +lean_dec(x_36); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_31); +lean_ctor_set(x_38, 1, x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_38); +lean_ctor_set(x_39, 1, x_32); +return x_39; +} +else +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_array_fget(x_13, x_33); +lean_dec(x_13); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_31); +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_32); +return x_43; +} +} +} +else +{ +uint8_t x_44; +lean_dec(x_15); +lean_dec(x_13); +x_44 = !lean_is_exclusive(x_19); +if (x_44 == 0) +{ +return x_19; +} +else +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_19, 0); +x_46 = lean_ctor_get(x_19, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_19); +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_dec(x_13); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_48 = l_Lean_Meta_Simp_mkSimpDiagSummary___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_7); +return x_49; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___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; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__5___rarg(x_1, x_2, x_6, x_7, x_5); +lean_dec(x_2); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___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_object* x_6) { +_start: +{ +lean_object* x_7; +x_7 = l_Lean_PersistentHashMap_foldlMAux_traverse___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__6___rarg(x_1, x_2, x_3, x_4, x_5, x_6); +lean_dec(x_3); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Array_qsort_sort___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__7(x_1, x_2, x_3, x_4); +lean_dec(x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___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_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +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_Meta_Simp_mkSimpDiagSummary___spec__8(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_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_Meta_Simp_mkSimpDiagSummary___lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("simp", 4); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_Meta_Simp_reportDiag___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("used theorems", 13); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tried theorems", 14); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__5() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("tried congruence theorems", 25); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__6() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("use `set_option diagnostics.threshold ` to control threshold for reporting counters", 88); +return x_1; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Simp_reportDiag___closed__6; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_Meta_Simp_reportDiag___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Meta_Simp_reportDiag___closed__7; +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_Meta_Simp_reportDiag(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_7 = l_Lean_isDiagnosticsEnabled(x_4, x_5, x_6); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_unbox(x_8); +lean_dec(x_8); +if (x_9 == 0) +{ +uint8_t x_10; +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_dec(x_11); +x_12 = lean_box(0); +lean_ctor_set(x_7, 0, x_12); +return x_7; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_7, 1); +lean_inc(x_13); +lean_dec(x_7); +x_14 = lean_box(0); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_13); +return x_15; +} +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_ctor_get(x_7, 1); +lean_inc(x_16); +lean_dec(x_7); +x_17 = lean_ctor_get(x_1, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_1, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_1, 2); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_box(0); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_17); +x_21 = l_Lean_Meta_Simp_mkSimpDiagSummary(x_17, x_20, x_2, x_3, x_4, x_5, x_16); +if (lean_obj_tag(x_21) == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_17); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +x_25 = l_Lean_Meta_Simp_mkSimpDiagSummary(x_18, x_24, x_2, x_3, x_4, x_5, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +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_Meta_Simp_mkSimpDiagSummary___closed__2; +x_29 = l_Lean_Meta_mkDiagSummary(x_19, x_28, x_2, x_3, x_4, x_5, x_27); +if (lean_obj_tag(x_29) == 0) +{ +uint8_t x_30; +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; +x_31 = lean_ctor_get(x_29, 0); +x_32 = lean_ctor_get(x_29, 1); +x_33 = l_Lean_Meta_DiagSummary_isEmpty(x_22); +if (x_33 == 0) +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; +lean_free_object(x_29); +x_34 = l_Lean_MessageData_nil; +x_35 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_36 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_37 = l_Lean_Meta_appendSection(x_34, x_35, x_36, x_22); +x_38 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_39 = l_Lean_Meta_appendSection(x_37, x_35, x_38, x_26); +x_40 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_41 = l_Lean_Meta_appendSection(x_39, x_35, x_40, x_31); +x_42 = l_Lean_Meta_Simp_reportDiag___closed__8; +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 = 0; +x_45 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_43, x_44, x_2, x_3, x_4, x_5, x_32); +return x_45; +} +else +{ +uint8_t x_46; +x_46 = l_Lean_Meta_DiagSummary_isEmpty(x_26); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; +lean_free_object(x_29); +x_47 = l_Lean_MessageData_nil; +x_48 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_49 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_50 = l_Lean_Meta_appendSection(x_47, x_48, x_49, x_22); +x_51 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_52 = l_Lean_Meta_appendSection(x_50, x_48, x_51, x_26); +x_53 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_54 = l_Lean_Meta_appendSection(x_52, x_48, x_53, x_31); +x_55 = l_Lean_Meta_Simp_reportDiag___closed__8; +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 = 0; +x_58 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_56, x_57, x_2, x_3, x_4, x_5, x_32); +return x_58; +} +else +{ +uint8_t x_59; +x_59 = l_Lean_Meta_DiagSummary_isEmpty(x_31); +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; uint8_t x_70; lean_object* x_71; +lean_free_object(x_29); +x_60 = l_Lean_MessageData_nil; +x_61 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_62 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_63 = l_Lean_Meta_appendSection(x_60, x_61, x_62, x_22); +x_64 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_65 = l_Lean_Meta_appendSection(x_63, x_61, x_64, x_26); +x_66 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_67 = l_Lean_Meta_appendSection(x_65, x_61, x_66, x_31); +x_68 = l_Lean_Meta_Simp_reportDiag___closed__8; +x_69 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_69, 0, x_67); +lean_ctor_set(x_69, 1, x_68); +x_70 = 0; +x_71 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_69, x_70, x_2, x_3, x_4, x_5, x_32); +return x_71; +} +else +{ +lean_object* x_72; +lean_dec(x_31); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_72 = lean_box(0); +lean_ctor_set(x_29, 0, x_72); +return x_29; +} +} +} +} +else +{ +lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_73 = lean_ctor_get(x_29, 0); +x_74 = lean_ctor_get(x_29, 1); +lean_inc(x_74); +lean_inc(x_73); +lean_dec(x_29); +x_75 = l_Lean_Meta_DiagSummary_isEmpty(x_22); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; +x_76 = l_Lean_MessageData_nil; +x_77 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_78 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_79 = l_Lean_Meta_appendSection(x_76, x_77, x_78, x_22); +x_80 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_81 = l_Lean_Meta_appendSection(x_79, x_77, x_80, x_26); +x_82 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_83 = l_Lean_Meta_appendSection(x_81, x_77, x_82, x_73); +x_84 = l_Lean_Meta_Simp_reportDiag___closed__8; +x_85 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_85, 0, x_83); +lean_ctor_set(x_85, 1, x_84); +x_86 = 0; +x_87 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_85, x_86, x_2, x_3, x_4, x_5, x_74); +return x_87; +} +else +{ +uint8_t x_88; +x_88 = l_Lean_Meta_DiagSummary_isEmpty(x_26); +if (x_88 == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; +x_89 = l_Lean_MessageData_nil; +x_90 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_91 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_92 = l_Lean_Meta_appendSection(x_89, x_90, x_91, x_22); +x_93 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_94 = l_Lean_Meta_appendSection(x_92, x_90, x_93, x_26); +x_95 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_96 = l_Lean_Meta_appendSection(x_94, x_90, x_95, x_73); +x_97 = l_Lean_Meta_Simp_reportDiag___closed__8; +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +x_99 = 0; +x_100 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_98, x_99, x_2, x_3, x_4, x_5, x_74); +return x_100; +} +else +{ +uint8_t x_101; +x_101 = l_Lean_Meta_DiagSummary_isEmpty(x_73); +if (x_101 == 0) +{ +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; +x_102 = l_Lean_MessageData_nil; +x_103 = l_Lean_Meta_Simp_reportDiag___closed__2; +x_104 = l_Lean_Meta_Simp_reportDiag___closed__3; +x_105 = l_Lean_Meta_appendSection(x_102, x_103, x_104, x_22); +x_106 = l_Lean_Meta_Simp_reportDiag___closed__4; +x_107 = l_Lean_Meta_appendSection(x_105, x_103, x_106, x_26); +x_108 = l_Lean_Meta_Simp_reportDiag___closed__5; +x_109 = l_Lean_Meta_appendSection(x_107, x_103, x_108, x_73); +x_110 = l_Lean_Meta_Simp_reportDiag___closed__8; +x_111 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +x_112 = 0; +x_113 = l_Lean_log___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__6(x_111, x_112, x_2, x_3, x_4, x_5, x_74); +return x_113; +} +else +{ +lean_object* x_114; lean_object* x_115; +lean_dec(x_73); +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_114 = lean_box(0); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_114); +lean_ctor_set(x_115, 1, x_74); +return x_115; +} +} +} +} +} +else +{ +uint8_t x_116; +lean_dec(x_26); +lean_dec(x_22); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_116 = !lean_is_exclusive(x_29); +if (x_116 == 0) +{ +return x_29; +} +else +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_117 = lean_ctor_get(x_29, 0); +x_118 = lean_ctor_get(x_29, 1); +lean_inc(x_118); +lean_inc(x_117); +lean_dec(x_29); +x_119 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_119, 0, x_117); +lean_ctor_set(x_119, 1, x_118); +return x_119; +} +} +} +else +{ +uint8_t x_120; +lean_dec(x_22); +lean_dec(x_19); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_120 = !lean_is_exclusive(x_25); +if (x_120 == 0) +{ +return x_25; +} +else +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_121 = lean_ctor_get(x_25, 0); +x_122 = lean_ctor_get(x_25, 1); +lean_inc(x_122); +lean_inc(x_121); +lean_dec(x_25); +x_123 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +return x_123; +} +} +} +else +{ +uint8_t x_124; +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_124 = !lean_is_exclusive(x_21); +if (x_124 == 0) +{ +return x_21; +} +else +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_21, 0); +x_126 = lean_ctor_get(x_21, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_21); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; +} +} +} +} +} +lean_object* initialize_Lean_Meta_Diagnostics(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Simp_Types(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Simp_Diagnostics(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_Meta_Diagnostics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Types(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1 = _init_l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_collectAboveThreshold___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__6); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__7); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__8); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__9); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__10); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__11); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__1___closed__12); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__2); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___lambda__2___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__1); +l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_mkSimpDiagSummary___spec__8___closed__2); +l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1 = _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_mkSimpDiagSummary___closed__1); +l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2 = _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_mkSimpDiagSummary___closed__2); +l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3 = _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_mkSimpDiagSummary___closed__3); +l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4 = _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_mkSimpDiagSummary___closed__4); +l_Lean_Meta_Simp_mkSimpDiagSummary___closed__5 = _init_l_Lean_Meta_Simp_mkSimpDiagSummary___closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_mkSimpDiagSummary___closed__5); +l_Lean_Meta_Simp_reportDiag___closed__1 = _init_l_Lean_Meta_Simp_reportDiag___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__1); +l_Lean_Meta_Simp_reportDiag___closed__2 = _init_l_Lean_Meta_Simp_reportDiag___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__2); +l_Lean_Meta_Simp_reportDiag___closed__3 = _init_l_Lean_Meta_Simp_reportDiag___closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__3); +l_Lean_Meta_Simp_reportDiag___closed__4 = _init_l_Lean_Meta_Simp_reportDiag___closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__4); +l_Lean_Meta_Simp_reportDiag___closed__5 = _init_l_Lean_Meta_Simp_reportDiag___closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__5); +l_Lean_Meta_Simp_reportDiag___closed__6 = _init_l_Lean_Meta_Simp_reportDiag___closed__6(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__6); +l_Lean_Meta_Simp_reportDiag___closed__7 = _init_l_Lean_Meta_Simp_reportDiag___closed__7(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__7); +l_Lean_Meta_Simp_reportDiag___closed__8 = _init_l_Lean_Meta_Simp_reportDiag___closed__8(); +lean_mark_persistent(l_Lean_Meta_Simp_reportDiag___closed__8); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c index 053d3816b89d..e58f1761ef54 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Main.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.Simp.Main -// Imports: Lean.Meta.Transform Lean.Meta.Tactic.Replace Lean.Meta.Tactic.UnifyEq Lean.Meta.Tactic.Simp.Rewrite Lean.Meta.Match.Value +// Imports: Lean.Meta.Transform Lean.Meta.Tactic.Replace Lean.Meta.Tactic.UnifyEq Lean.Meta.Tactic.Simp.Rewrite Lean.Meta.Tactic.Simp.Diagnostics Lean.Meta.Match.Value #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -29,6 +29,7 @@ lean_object* l_Lean_Meta_reduceRecMatcher_x3f(lean_object*, lean_object*, lean_o static lean_object* l_Lean_Meta_Simp_simpLambda___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_lambdaTelescopeDSimp___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_Meta_Simp_cacheResult___boxed(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_profileitM___at_Lean_Meta_synthInstance_x3f___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_Meta_Simp_simpForall___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_dsimpGoal___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_HashMapImp_expand___at_Lean_Meta_Simp_cacheResult___spec__8(lean_object*, lean_object*); @@ -60,6 +61,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simpGoal___lambda__1(lean_object*, lean_obj LEAN_EXPORT lean_object* l_Lean_Meta_simpTarget___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_Meta_dsimpGoal(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_dsimp(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_main_simpMain(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_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*); @@ -101,6 +103,7 @@ lean_object* l_Lean_MVarId_checkNotAssigned(lean_object*, lean_object*, lean_obj static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_applySimpResultToProp___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__2(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_main_simpMain___closed__1; 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*); static lean_object* l_Lean_Meta_Simp_isCharLit___closed__1; @@ -129,6 +132,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tact LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLet___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimpLetCase(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_Meta_Simp_main___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*); 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*); @@ -190,6 +194,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLet___at___private_Lea LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___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_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Simp_simpLet___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpContext(lean_object*); lean_object* l_Lean_Meta_throwNestedTacticEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_Simp_cacheResult___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -249,7 +254,6 @@ extern lean_object* l_Lean_Meta_smartUnfolding; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Meta_Simp_simpProj___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__5___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isHeadBetaTargetFn(uint8_t, lean_object*); -static lean_object* l_Lean_Meta_Simp_dsimpMain___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpStep(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_EXPORT lean_object* l_Lean_Meta_dsimpGoal___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*); @@ -276,6 +280,7 @@ lean_object* l_Lean_Meta_mkEqMP(lean_object*, lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpForall___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_Meta_mkLetValCongr(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_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__5; static lean_object* l_Lean_getProjectionFnInfo_x3f___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceProjFn_x3f___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__6___boxed(lean_object**); LEAN_EXPORT uint8_t l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_isMatchDef_go(lean_object*, lean_object*); @@ -284,7 +289,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_foldRawNatLit___boxed(lean_object*, le static lean_object* l_Lean_Meta_Simp_isOfNatNatLit___closed__2; static lean_object* l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_processCongrHypothesis___spec__1___lambda__3___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__1; -static lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__2___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isOfNatNatLit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpImpl_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -295,7 +299,6 @@ 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*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimpLetCase_toCtorIdx(uint8_t); -static lean_object* l_Lean_Meta_Simp_dsimpMain___closed__2; lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_InlineAttrs_0__Lean_Compiler_isValidMacroInline___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__5___closed__2; @@ -311,6 +314,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at___private_ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_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_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__4___lambda__1___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimpLetCase_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_dsimpGoal___lambda__2___closed__2; lean_object* l_Lean_instantiateMVars___at___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___spec__7(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___lambda__1___closed__2; @@ -343,14 +347,12 @@ lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_Simp_tryTheoremWithExtr lean_object* l_Lean_Expr_abstractM(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_Simp_simpProj___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_visitFn___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___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_EXPORT lean_object* l_Lean_Meta_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3; static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__5___closed__4; static lean_object* l_Lean_Meta_Simp_simpLet___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__5___lambda__1___boxed(lean_object**); -lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpArrow___lambda__4___closed__1; uint8_t l_Lean_Expr_isMVar(lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -384,7 +386,6 @@ static lean_object* l_Lean_Meta_Simp_simpForall___lambda__5___closed__1; static lean_object* l_Lean_Meta_Simp_simpArrow___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_unfold_x3f(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_simpForall___lambda__5___closed__8; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Meta_reduceMatcher_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__16___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -439,6 +440,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Sim lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isArrow(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_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLoop_visitPostContinue(lean_object*, lean_object*, 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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_dsimp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -504,13 +506,14 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_red lean_object* l_Lean_LocalDecl_userName(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduce___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_Simp_SimpLetCase_toCtorIdx___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_simp___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___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_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Main___hyg_5_(lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpForall___closed__2; +lean_object* l_Lean_Meta_Simp_recordCongrTheorem___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_Simp_simpLet___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimpMain_dsimpMain(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFalse(lean_object*); static lean_object* l_Lean_Meta_Simp_withNewLemmas___rarg___closed__4; lean_object* l_Lean_Meta_isConstructorApp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -571,6 +574,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLet___lambda__2(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Meta_simpGoal___lambda__2___boxed(lean_object**); static lean_object* l_Lean_Meta_Simp_simpProj___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_congr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_reportDiag(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_unfoldDefinition_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpProj___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkForallFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -580,6 +584,7 @@ static lean_object* l_Lean_Meta_Simp_simpArrow___closed__2; lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_exprDependsOn___at_Lean_Meta_Simp_simpProj___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__10___boxed(lean_object**); +static lean_object* l_Lean_Meta_dsimp___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLoop_visitPostContinue___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_main___lambda__1___boxed(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*); @@ -636,8 +641,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withNewLemmas___rarg(lean_object*, lea static lean_object* l_Lean_Meta_simpTargetCore___closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_simpImpl_go___spec__6(lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___closed__1; lean_object* l_Lean_Expr_letFunAppArgs_x3f(lean_object*); -lean_object* l_Lean_Meta_Simp_mkMethods(lean_object*, lean_object*); +lean_object* l_Lean_Meta_Simp_mkMethods(lean_object*, lean_object*, uint8_t); static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_processCongrHypothesis___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_PersistentArray_anyM___at_Lean_Meta_Simp_simpProj___spec__5___boxed(lean_object*, lean_object*); @@ -649,7 +655,7 @@ lean_object* l_Lean_Meta_mkCongrFun(lean_object*, lean_object*, lean_object*, le 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_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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_Lean_SMap_insert___at_Lean_Meta_Simp_cacheResult___spec__1(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*); @@ -665,6 +671,7 @@ uint8_t l_Lean_Expr_hasFVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_simpLocalDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpForall___lambda__5___closed__3; lean_object* l_Lean_Meta_Simp_Result_mkEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_processCongrHypothesis___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_Expr_betaRev(lean_object*, lean_object*, uint8_t, uint8_t); uint8_t l_Lean_Meta_SimpTheoremsArray_isLetDeclToUnfold(lean_object*, lean_object*); @@ -676,7 +683,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_anyM___at_Lean_Meta_Simp_simpPro 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*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___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_EXPORT lean_object* l_Lean_Meta_simpLocalDecl(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_PersistentHashMap_findAux___at_Lean_Meta_Simp_simpImpl_go___spec__3(lean_object*, size_t, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); @@ -687,6 +693,7 @@ LEAN_EXPORT lean_object* l_Lean_isProjectionFn___at___private_Lean_Meta_Tactic_S static lean_object* l_Lean_Meta_Simp_isOfNatNatLit___closed__1; 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* lean_local_ctx_num_indices(lean_object*); lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_throwCongrHypothesisFailed___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_dsimpImpl___spec__8___boxed(lean_object*, lean_object*); @@ -728,10 +735,9 @@ 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; -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*); lean_object* lean_nat_add(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l_Lean_Meta_Simp_isCharLit___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7___boxed(lean_object**); @@ -756,6 +762,7 @@ lean_object* l_Lean_Meta_reduceProj_x3f(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(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_SimpTheoremsArray_addTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpLoop___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpContext___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_main___closed__2; lean_object* l_Lean_Meta_mkImpCongrCtx(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_consumeMData(lean_object*); @@ -787,14 +794,13 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_red LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpImpl___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_simpArrow___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___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); -static lean_object* l_Lean_Meta_Simp_main___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_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; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpLet___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimpMain___boxed(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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Main___hyg_5____closed__1() { _start: { @@ -8191,235 +8197,581 @@ return x_4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withNewLemmas___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_36; lean_object* x_37; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_11 = lean_st_ref_get(x_5, x_10); +lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_11 = l_Lean_Meta_Simp_getConfig___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_10); x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 0); -lean_inc(x_14); +x_13 = lean_ctor_get_uint8(x_12, sizeof(void*)*2); lean_dec(x_12); -x_57 = lean_st_ref_take(x_5, x_13); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); +if (x_13 == 0) +{ +uint8_t x_14; +x_14 = lean_ctor_get_uint8(x_3, sizeof(void*)*5); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_15 = lean_ctor_get(x_11, 1); +lean_inc(x_15); +lean_dec(x_11); +x_16 = lean_st_ref_get(x_5, 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_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_st_ref_take(x_5, 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 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_24 = lean_ctor_get(x_21, 0); +lean_dec(x_24); +x_25 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +lean_ctor_set(x_21, 0, x_25); +x_26 = lean_st_ref_set(x_5, x_21, x_22); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +lean_inc(x_5); +x_28 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_27); +if (lean_obj_tag(x_28) == 0) +{ +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +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_5, 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; uint8_t x_37; +x_35 = lean_ctor_get(x_32, 0); +lean_dec(x_35); +lean_ctor_set(x_32, 0, x_19); +x_36 = lean_st_ref_set(x_5, x_32, x_33); +lean_dec(x_5); +x_37 = !lean_is_exclusive(x_36); +if (x_37 == 0) +{ +lean_object* x_38; +x_38 = lean_ctor_get(x_36, 0); +lean_dec(x_38); +lean_ctor_set(x_36, 0, x_29); +return x_36; +} +else +{ +lean_object* x_39; lean_object* x_40; +x_39 = lean_ctor_get(x_36, 1); +lean_inc(x_39); +lean_dec(x_36); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_29); +lean_ctor_set(x_40, 1, x_39); +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; +x_41 = lean_ctor_get(x_32, 1); +x_42 = lean_ctor_get(x_32, 2); +x_43 = lean_ctor_get(x_32, 3); +x_44 = lean_ctor_get(x_32, 4); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_32); +x_45 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_45, 0, x_19); +lean_ctor_set(x_45, 1, x_41); +lean_ctor_set(x_45, 2, x_42); +lean_ctor_set(x_45, 3, x_43); +lean_ctor_set(x_45, 4, x_44); +x_46 = lean_st_ref_set(x_5, x_45, x_33); +lean_dec(x_5); +x_47 = lean_ctor_get(x_46, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_46)) { + lean_ctor_release(x_46, 0); + lean_ctor_release(x_46, 1); + x_48 = x_46; +} else { + lean_dec_ref(x_46); + x_48 = lean_box(0); +} +if (lean_is_scalar(x_48)) { + x_49 = lean_alloc_ctor(0, 2, 0); +} else { + x_49 = x_48; +} +lean_ctor_set(x_49, 0, x_29); +lean_ctor_set(x_49, 1, x_47); +return x_49; +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_50 = lean_ctor_get(x_28, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_28, 1); +lean_inc(x_51); +lean_dec(x_28); +x_52 = lean_st_ref_take(x_5, x_51); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = !lean_is_exclusive(x_53); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; uint8_t x_58; +x_56 = lean_ctor_get(x_53, 0); +lean_dec(x_56); +lean_ctor_set(x_53, 0, x_19); +x_57 = lean_st_ref_set(x_5, x_53, x_54); +lean_dec(x_5); +x_58 = !lean_is_exclusive(x_57); +if (x_58 == 0) +{ +lean_object* x_59; +x_59 = lean_ctor_get(x_57, 0); +lean_dec(x_59); +lean_ctor_set_tag(x_57, 1); +lean_ctor_set(x_57, 0, x_50); +return x_57; +} +else +{ +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_57, 1); +lean_inc(x_60); lean_dec(x_57); -x_60 = !lean_is_exclusive(x_58); -if (x_60 == 0) +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_50); +lean_ctor_set(x_61, 1, x_60); +return x_61; +} +} +else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_61 = lean_ctor_get(x_58, 0); -lean_dec(x_61); -x_62 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; -lean_ctor_set(x_58, 0, x_62); -x_63 = lean_st_ref_set(x_5, x_58, x_59); -x_64 = lean_ctor_get(x_63, 1); +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_62 = lean_ctor_get(x_53, 1); +x_63 = lean_ctor_get(x_53, 2); +x_64 = lean_ctor_get(x_53, 3); +x_65 = lean_ctor_get(x_53, 4); +lean_inc(x_65); lean_inc(x_64); -lean_dec(x_63); -x_65 = l_Lean_Meta_Simp_getConfig___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_64); -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get_uint8(x_66, sizeof(void*)*2); -lean_dec(x_66); -if (x_67 == 0) -{ -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_65, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_53); +x_66 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_66, 0, x_19); +lean_ctor_set(x_66, 1, x_62); +lean_ctor_set(x_66, 2, x_63); +lean_ctor_set(x_66, 3, x_64); +lean_ctor_set(x_66, 4, x_65); +x_67 = lean_st_ref_set(x_5, x_66, x_54); +lean_dec(x_5); +x_68 = lean_ctor_get(x_67, 1); lean_inc(x_68); -lean_dec(x_65); -lean_inc(x_5); -x_69 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_68); -if (lean_obj_tag(x_69) == 0) +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_69 = x_67; +} else { + lean_dec_ref(x_67); + x_69 = lean_box(0); +} +if (lean_is_scalar(x_69)) { + x_70 = lean_alloc_ctor(1, 2, 0); +} else { + x_70 = x_69; + lean_ctor_set_tag(x_70, 1); +} +lean_ctor_set(x_70, 0, x_50); +lean_ctor_set(x_70, 1, x_68); +return x_70; +} +} +} +else { -lean_object* x_70; lean_object* x_71; -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); +lean_object* x_71; lean_object* x_72; lean_object* x_73; 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_71 = lean_ctor_get(x_21, 1); +x_72 = lean_ctor_get(x_21, 2); +x_73 = lean_ctor_get(x_21, 3); +x_74 = lean_ctor_get(x_21, 4); +lean_inc(x_74); +lean_inc(x_73); +lean_inc(x_72); lean_inc(x_71); -lean_dec(x_69); -x_15 = x_70; -x_16 = x_71; -goto block_35; +lean_dec(x_21); +x_75 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_76 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_71); +lean_ctor_set(x_76, 2, x_72); +lean_ctor_set(x_76, 3, x_73); +lean_ctor_set(x_76, 4, x_74); +x_77 = lean_st_ref_set(x_5, x_76, x_22); +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +lean_inc(x_5); +x_79 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_78); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; 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_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_5, 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, 1); +lean_inc(x_85); +x_86 = lean_ctor_get(x_83, 2); +lean_inc(x_86); +x_87 = lean_ctor_get(x_83, 3); +lean_inc(x_87); +x_88 = lean_ctor_get(x_83, 4); +lean_inc(x_88); +if (lean_is_exclusive(x_83)) { + lean_ctor_release(x_83, 0); + lean_ctor_release(x_83, 1); + lean_ctor_release(x_83, 2); + lean_ctor_release(x_83, 3); + lean_ctor_release(x_83, 4); + x_89 = x_83; +} else { + lean_dec_ref(x_83); + x_89 = lean_box(0); +} +if (lean_is_scalar(x_89)) { + x_90 = lean_alloc_ctor(0, 5, 0); +} else { + x_90 = x_89; +} +lean_ctor_set(x_90, 0, x_19); +lean_ctor_set(x_90, 1, x_85); +lean_ctor_set(x_90, 2, x_86); +lean_ctor_set(x_90, 3, x_87); +lean_ctor_set(x_90, 4, x_88); +x_91 = lean_st_ref_set(x_5, x_90, x_84); +lean_dec(x_5); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +if (lean_is_exclusive(x_91)) { + lean_ctor_release(x_91, 0); + lean_ctor_release(x_91, 1); + x_93 = x_91; +} else { + lean_dec_ref(x_91); + x_93 = lean_box(0); +} +if (lean_is_scalar(x_93)) { + x_94 = lean_alloc_ctor(0, 2, 0); +} else { + x_94 = x_93; +} +lean_ctor_set(x_94, 0, x_80); +lean_ctor_set(x_94, 1, x_92); +return x_94; } else { -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_69, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_69, 1); -lean_inc(x_73); -lean_dec(x_69); -x_36 = x_72; -x_37 = x_73; -goto block_56; +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_95 = lean_ctor_get(x_79, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_79, 1); +lean_inc(x_96); +lean_dec(x_79); +x_97 = lean_st_ref_take(x_5, x_96); +x_98 = lean_ctor_get(x_97, 0); +lean_inc(x_98); +x_99 = lean_ctor_get(x_97, 1); +lean_inc(x_99); +lean_dec(x_97); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +x_101 = lean_ctor_get(x_98, 2); +lean_inc(x_101); +x_102 = lean_ctor_get(x_98, 3); +lean_inc(x_102); +x_103 = lean_ctor_get(x_98, 4); +lean_inc(x_103); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + lean_ctor_release(x_98, 2); + lean_ctor_release(x_98, 3); + lean_ctor_release(x_98, 4); + x_104 = x_98; +} else { + lean_dec_ref(x_98); + x_104 = lean_box(0); +} +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(0, 5, 0); +} else { + x_105 = x_104; +} +lean_ctor_set(x_105, 0, x_19); +lean_ctor_set(x_105, 1, x_100); +lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 3, x_102); +lean_ctor_set(x_105, 4, x_103); +x_106 = lean_st_ref_set(x_5, x_105, x_99); +lean_dec(x_5); +x_107 = lean_ctor_get(x_106, 1); +lean_inc(x_107); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_108 = x_106; +} else { + lean_dec_ref(x_106); + x_108 = lean_box(0); +} +if (lean_is_scalar(x_108)) { + x_109 = lean_alloc_ctor(1, 2, 0); +} else { + x_109 = x_108; + lean_ctor_set_tag(x_109, 1); +} +lean_ctor_set(x_109, 0, x_95); +lean_ctor_set(x_109, 1, x_107); +return x_109; +} } } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; size_t x_82; size_t x_83; lean_object* x_84; -x_74 = lean_ctor_get(x_65, 1); -lean_inc(x_74); -lean_dec(x_65); -x_75 = l_Lean_Meta_Simp_getSimpTheorems___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_74); -x_76 = lean_ctor_get(x_75, 0); -lean_inc(x_76); -x_77 = lean_ctor_get(x_75, 1); -lean_inc(x_77); -lean_dec(x_75); -x_78 = 0; -x_79 = lean_box(x_78); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_76); -lean_ctor_set(x_80, 1, x_79); -x_81 = lean_array_get_size(x_1); -x_82 = lean_usize_of_nat(x_81); -lean_dec(x_81); -x_83 = 0; +lean_object* x_110; lean_object* x_111; +x_110 = lean_ctor_get(x_11, 1); +lean_inc(x_110); +lean_dec(x_11); +x_111 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_110); +return x_111; +} +} +else +{ +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_139; lean_object* x_140; lean_object* x_161; lean_object* x_162; lean_object* x_163; uint8_t x_164; +x_112 = lean_ctor_get(x_11, 1); +lean_inc(x_112); +lean_dec(x_11); +x_113 = lean_st_ref_get(x_5, x_112); +x_114 = lean_ctor_get(x_113, 0); +lean_inc(x_114); +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_ctor_get(x_114, 0); +lean_inc(x_116); +lean_dec(x_114); +x_161 = lean_st_ref_take(x_5, x_115); +x_162 = lean_ctor_get(x_161, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_161, 1); +lean_inc(x_163); +lean_dec(x_161); +x_164 = !lean_is_exclusive(x_162); +if (x_164 == 0) +{ +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; size_t x_176; size_t x_177; lean_object* x_178; +x_165 = lean_ctor_get(x_162, 0); +lean_dec(x_165); +x_166 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +lean_ctor_set(x_162, 0, x_166); +x_167 = lean_st_ref_set(x_5, x_162, x_163); +x_168 = lean_ctor_get(x_167, 1); +lean_inc(x_168); +lean_dec(x_167); +x_169 = l_Lean_Meta_Simp_getSimpTheorems___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_168); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +lean_dec(x_169); +x_172 = 0; +x_173 = lean_box(x_172); +x_174 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_174, 0, x_170); +lean_ctor_set(x_174, 1, x_173); +x_175 = lean_array_get_size(x_1); +x_176 = lean_usize_of_nat(x_175); +lean_dec(x_175); +x_177 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_84 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_withNewLemmas___spec__1(x_1, x_82, x_83, x_80, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_77); -if (lean_obj_tag(x_84) == 0) +x_178 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_withNewLemmas___spec__1(x_1, x_176, x_177, x_174, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_171); +if (lean_obj_tag(x_178) == 0) { -lean_object* x_85; lean_object* x_86; uint8_t x_87; -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_85, 1); -lean_inc(x_86); -x_87 = lean_unbox(x_86); -lean_dec(x_86); -if (x_87 == 0) +lean_object* x_179; lean_object* x_180; uint8_t x_181; +x_179 = lean_ctor_get(x_178, 0); +lean_inc(x_179); +x_180 = lean_ctor_get(x_179, 1); +lean_inc(x_180); +x_181 = lean_unbox(x_180); +lean_dec(x_180); +if (x_181 == 0) { -lean_object* x_88; lean_object* x_89; -lean_dec(x_85); -x_88 = lean_ctor_get(x_84, 1); -lean_inc(x_88); -lean_dec(x_84); +lean_object* x_182; lean_object* x_183; +lean_dec(x_179); +x_182 = lean_ctor_get(x_178, 1); +lean_inc(x_182); +lean_dec(x_178); lean_inc(x_5); -x_89 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_88); -if (lean_obj_tag(x_89) == 0) +x_183 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_182); +if (lean_obj_tag(x_183) == 0) { -lean_object* x_90; lean_object* x_91; -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_15 = x_90; -x_16 = x_91; -goto block_35; +lean_object* x_184; lean_object* x_185; +x_184 = lean_ctor_get(x_183, 0); +lean_inc(x_184); +x_185 = lean_ctor_get(x_183, 1); +lean_inc(x_185); +lean_dec(x_183); +x_117 = x_184; +x_118 = x_185; +goto block_138; } else { -lean_object* x_92; lean_object* x_93; -x_92 = lean_ctor_get(x_89, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_89, 1); -lean_inc(x_93); -lean_dec(x_89); -x_36 = x_92; -x_37 = x_93; -goto block_56; +lean_object* x_186; lean_object* x_187; +x_186 = lean_ctor_get(x_183, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_183, 1); +lean_inc(x_187); +lean_dec(x_183); +x_139 = x_186; +x_140 = x_187; +goto block_160; } } else { -lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_94 = lean_ctor_get(x_84, 1); -lean_inc(x_94); -lean_dec(x_84); -x_95 = lean_ctor_get(x_85, 0); -lean_inc(x_95); -lean_dec(x_85); -x_96 = !lean_is_exclusive(x_4); -if (x_96 == 0) +lean_object* x_188; lean_object* x_189; uint8_t x_190; +x_188 = lean_ctor_get(x_178, 1); +lean_inc(x_188); +lean_dec(x_178); +x_189 = lean_ctor_get(x_179, 0); +lean_inc(x_189); +lean_dec(x_179); +x_190 = !lean_is_exclusive(x_4); +if (x_190 == 0) { -lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_4, 1); -lean_dec(x_97); -lean_ctor_set(x_4, 1, x_95); +lean_object* x_191; lean_object* x_192; +x_191 = lean_ctor_get(x_4, 1); +lean_dec(x_191); +lean_ctor_set(x_4, 1, x_189); lean_inc(x_5); -x_98 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_94); -if (lean_obj_tag(x_98) == 0) +x_192 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_188); +if (lean_obj_tag(x_192) == 0) { -lean_object* x_99; lean_object* x_100; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -lean_dec(x_98); -x_15 = x_99; -x_16 = x_100; -goto block_35; +lean_object* x_193; lean_object* x_194; +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_117 = x_193; +x_118 = x_194; +goto block_138; } else { -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_98, 1); -lean_inc(x_102); -lean_dec(x_98); -x_36 = x_101; -x_37 = x_102; -goto block_56; +lean_object* x_195; lean_object* x_196; +x_195 = lean_ctor_get(x_192, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_192, 1); +lean_inc(x_196); +lean_dec(x_192); +x_139 = x_195; +x_140 = x_196; +goto block_160; } } else { -lean_object* x_103; uint32_t x_104; lean_object* x_105; lean_object* x_106; uint32_t x_107; lean_object* x_108; lean_object* x_109; -x_103 = lean_ctor_get(x_4, 0); -x_104 = lean_ctor_get_uint32(x_4, sizeof(void*)*4); -x_105 = lean_ctor_get(x_4, 2); -x_106 = lean_ctor_get(x_4, 3); -x_107 = lean_ctor_get_uint32(x_4, sizeof(void*)*4 + 4); -lean_inc(x_106); -lean_inc(x_105); -lean_inc(x_103); +lean_object* x_197; uint32_t x_198; lean_object* x_199; lean_object* x_200; uint32_t x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_197 = lean_ctor_get(x_4, 0); +x_198 = lean_ctor_get_uint32(x_4, sizeof(void*)*5); +x_199 = lean_ctor_get(x_4, 2); +x_200 = lean_ctor_get(x_4, 3); +x_201 = lean_ctor_get_uint32(x_4, sizeof(void*)*5 + 4); +x_202 = lean_ctor_get(x_4, 4); +lean_inc(x_202); +lean_inc(x_200); +lean_inc(x_199); +lean_inc(x_197); lean_dec(x_4); -x_108 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_108, 0, x_103); -lean_ctor_set(x_108, 1, x_95); -lean_ctor_set(x_108, 2, x_105); -lean_ctor_set(x_108, 3, x_106); -lean_ctor_set_uint32(x_108, sizeof(void*)*4, x_104); -lean_ctor_set_uint32(x_108, sizeof(void*)*4 + 4, x_107); +x_203 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_203, 0, x_197); +lean_ctor_set(x_203, 1, x_189); +lean_ctor_set(x_203, 2, x_199); +lean_ctor_set(x_203, 3, x_200); +lean_ctor_set(x_203, 4, x_202); +lean_ctor_set_uint32(x_203, sizeof(void*)*5, x_198); +lean_ctor_set_uint32(x_203, sizeof(void*)*5 + 4, x_201); lean_inc(x_5); -x_109 = lean_apply_8(x_2, x_3, x_108, x_5, x_6, x_7, x_8, x_9, x_94); -if (lean_obj_tag(x_109) == 0) +x_204 = lean_apply_8(x_2, x_3, x_203, x_5, x_6, x_7, x_8, x_9, x_188); +if (lean_obj_tag(x_204) == 0) { -lean_object* x_110; lean_object* x_111; -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_15 = x_110; -x_16 = x_111; -goto block_35; +lean_object* x_205; lean_object* x_206; +x_205 = lean_ctor_get(x_204, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_204, 1); +lean_inc(x_206); +lean_dec(x_204); +x_117 = x_205; +x_118 = x_206; +goto block_138; } else { -lean_object* x_112; lean_object* x_113; -x_112 = lean_ctor_get(x_109, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_109, 1); -lean_inc(x_113); -lean_dec(x_109); -x_36 = x_112; -x_37 = x_113; -goto block_56; +lean_object* x_207; lean_object* x_208; +x_207 = lean_ctor_get(x_204, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_204, 1); +lean_inc(x_208); +lean_dec(x_204); +x_139 = x_207; +x_140 = x_208; +goto block_160; } } } } else { -lean_object* x_114; lean_object* x_115; +lean_object* x_209; lean_object* x_210; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8427,213 +8779,175 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_114 = lean_ctor_get(x_84, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_84, 1); -lean_inc(x_115); -lean_dec(x_84); -x_36 = x_114; -x_37 = x_115; -goto block_56; -} -} -} -else -{ -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t x_125; -x_116 = lean_ctor_get(x_58, 1); -x_117 = lean_ctor_get(x_58, 2); -x_118 = lean_ctor_get(x_58, 3); -lean_inc(x_118); -lean_inc(x_117); -lean_inc(x_116); -lean_dec(x_58); -x_119 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; -x_120 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_116); -lean_ctor_set(x_120, 2, x_117); -lean_ctor_set(x_120, 3, x_118); -x_121 = lean_st_ref_set(x_5, x_120, x_59); -x_122 = lean_ctor_get(x_121, 1); -lean_inc(x_122); -lean_dec(x_121); -x_123 = l_Lean_Meta_Simp_getConfig___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_122); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get_uint8(x_124, sizeof(void*)*2); -lean_dec(x_124); -if (x_125 == 0) -{ -lean_object* x_126; lean_object* x_127; -x_126 = lean_ctor_get(x_123, 1); -lean_inc(x_126); -lean_dec(x_123); -lean_inc(x_5); -x_127 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_126); -if (lean_obj_tag(x_127) == 0) -{ -lean_object* x_128; lean_object* x_129; -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_ctor_get(x_127, 1); -lean_inc(x_129); -lean_dec(x_127); -x_15 = x_128; -x_16 = x_129; -goto block_35; -} -else -{ -lean_object* x_130; lean_object* x_131; -x_130 = lean_ctor_get(x_127, 0); -lean_inc(x_130); -x_131 = lean_ctor_get(x_127, 1); -lean_inc(x_131); -lean_dec(x_127); -x_36 = x_130; -x_37 = x_131; -goto block_56; +x_209 = lean_ctor_get(x_178, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_178, 1); +lean_inc(x_210); +lean_dec(x_178); +x_139 = x_209; +x_140 = x_210; +goto block_160; } } else { -lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; uint8_t x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; size_t x_140; size_t x_141; lean_object* x_142; -x_132 = lean_ctor_get(x_123, 1); -lean_inc(x_132); -lean_dec(x_123); -x_133 = l_Lean_Meta_Simp_getSimpTheorems___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_132); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -lean_dec(x_133); -x_136 = 0; -x_137 = lean_box(x_136); -x_138 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_138, 0, x_134); -lean_ctor_set(x_138, 1, x_137); -x_139 = lean_array_get_size(x_1); -x_140 = lean_usize_of_nat(x_139); -lean_dec(x_139); -x_141 = 0; +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; uint8_t x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; size_t x_226; size_t x_227; lean_object* x_228; +x_211 = lean_ctor_get(x_162, 1); +x_212 = lean_ctor_get(x_162, 2); +x_213 = lean_ctor_get(x_162, 3); +x_214 = lean_ctor_get(x_162, 4); +lean_inc(x_214); +lean_inc(x_213); +lean_inc(x_212); +lean_inc(x_211); +lean_dec(x_162); +x_215 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_216 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_216, 0, x_215); +lean_ctor_set(x_216, 1, x_211); +lean_ctor_set(x_216, 2, x_212); +lean_ctor_set(x_216, 3, x_213); +lean_ctor_set(x_216, 4, x_214); +x_217 = lean_st_ref_set(x_5, x_216, x_163); +x_218 = lean_ctor_get(x_217, 1); +lean_inc(x_218); +lean_dec(x_217); +x_219 = l_Lean_Meta_Simp_getSimpTheorems___rarg(x_4, x_5, x_6, x_7, x_8, x_9, x_218); +x_220 = lean_ctor_get(x_219, 0); +lean_inc(x_220); +x_221 = lean_ctor_get(x_219, 1); +lean_inc(x_221); +lean_dec(x_219); +x_222 = 0; +x_223 = lean_box(x_222); +x_224 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_224, 0, x_220); +lean_ctor_set(x_224, 1, x_223); +x_225 = lean_array_get_size(x_1); +x_226 = lean_usize_of_nat(x_225); +lean_dec(x_225); +x_227 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_142 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_withNewLemmas___spec__1(x_1, x_140, x_141, x_138, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_135); -if (lean_obj_tag(x_142) == 0) +x_228 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_withNewLemmas___spec__1(x_1, x_226, x_227, x_224, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_221); +if (lean_obj_tag(x_228) == 0) { -lean_object* x_143; lean_object* x_144; uint8_t x_145; -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_143, 1); -lean_inc(x_144); -x_145 = lean_unbox(x_144); -lean_dec(x_144); -if (x_145 == 0) +lean_object* x_229; lean_object* x_230; uint8_t x_231; +x_229 = lean_ctor_get(x_228, 0); +lean_inc(x_229); +x_230 = lean_ctor_get(x_229, 1); +lean_inc(x_230); +x_231 = lean_unbox(x_230); +lean_dec(x_230); +if (x_231 == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_143); -x_146 = lean_ctor_get(x_142, 1); -lean_inc(x_146); -lean_dec(x_142); +lean_object* x_232; lean_object* x_233; +lean_dec(x_229); +x_232 = lean_ctor_get(x_228, 1); +lean_inc(x_232); +lean_dec(x_228); lean_inc(x_5); -x_147 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_146); -if (lean_obj_tag(x_147) == 0) +x_233 = lean_apply_8(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_232); +if (lean_obj_tag(x_233) == 0) { -lean_object* x_148; lean_object* x_149; -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_15 = x_148; -x_16 = x_149; -goto block_35; +lean_object* x_234; lean_object* x_235; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_ctor_get(x_233, 1); +lean_inc(x_235); +lean_dec(x_233); +x_117 = x_234; +x_118 = x_235; +goto block_138; } else { -lean_object* x_150; lean_object* x_151; -x_150 = lean_ctor_get(x_147, 0); -lean_inc(x_150); -x_151 = lean_ctor_get(x_147, 1); -lean_inc(x_151); -lean_dec(x_147); -x_36 = x_150; -x_37 = x_151; -goto block_56; +lean_object* x_236; lean_object* x_237; +x_236 = lean_ctor_get(x_233, 0); +lean_inc(x_236); +x_237 = lean_ctor_get(x_233, 1); +lean_inc(x_237); +lean_dec(x_233); +x_139 = x_236; +x_140 = x_237; +goto block_160; } } else { -lean_object* x_152; lean_object* x_153; lean_object* x_154; uint32_t x_155; lean_object* x_156; lean_object* x_157; uint32_t x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_152 = lean_ctor_get(x_142, 1); -lean_inc(x_152); -lean_dec(x_142); -x_153 = lean_ctor_get(x_143, 0); -lean_inc(x_153); -lean_dec(x_143); -x_154 = lean_ctor_get(x_4, 0); -lean_inc(x_154); -x_155 = lean_ctor_get_uint32(x_4, sizeof(void*)*4); -x_156 = lean_ctor_get(x_4, 2); -lean_inc(x_156); -x_157 = lean_ctor_get(x_4, 3); -lean_inc(x_157); -x_158 = lean_ctor_get_uint32(x_4, sizeof(void*)*4 + 4); +lean_object* x_238; lean_object* x_239; lean_object* x_240; uint32_t x_241; lean_object* x_242; lean_object* x_243; uint32_t x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; +x_238 = lean_ctor_get(x_228, 1); +lean_inc(x_238); +lean_dec(x_228); +x_239 = lean_ctor_get(x_229, 0); +lean_inc(x_239); +lean_dec(x_229); +x_240 = lean_ctor_get(x_4, 0); +lean_inc(x_240); +x_241 = lean_ctor_get_uint32(x_4, sizeof(void*)*5); +x_242 = lean_ctor_get(x_4, 2); +lean_inc(x_242); +x_243 = lean_ctor_get(x_4, 3); +lean_inc(x_243); +x_244 = lean_ctor_get_uint32(x_4, sizeof(void*)*5 + 4); +x_245 = lean_ctor_get(x_4, 4); +lean_inc(x_245); if (lean_is_exclusive(x_4)) { lean_ctor_release(x_4, 0); lean_ctor_release(x_4, 1); lean_ctor_release(x_4, 2); lean_ctor_release(x_4, 3); - x_159 = x_4; + lean_ctor_release(x_4, 4); + x_246 = x_4; } else { lean_dec_ref(x_4); - x_159 = lean_box(0); + x_246 = lean_box(0); } -if (lean_is_scalar(x_159)) { - x_160 = lean_alloc_ctor(0, 4, 8); +if (lean_is_scalar(x_246)) { + x_247 = lean_alloc_ctor(0, 5, 8); } else { - x_160 = x_159; + x_247 = x_246; } -lean_ctor_set(x_160, 0, x_154); -lean_ctor_set(x_160, 1, x_153); -lean_ctor_set(x_160, 2, x_156); -lean_ctor_set(x_160, 3, x_157); -lean_ctor_set_uint32(x_160, sizeof(void*)*4, x_155); -lean_ctor_set_uint32(x_160, sizeof(void*)*4 + 4, x_158); +lean_ctor_set(x_247, 0, x_240); +lean_ctor_set(x_247, 1, x_239); +lean_ctor_set(x_247, 2, x_242); +lean_ctor_set(x_247, 3, x_243); +lean_ctor_set(x_247, 4, x_245); +lean_ctor_set_uint32(x_247, sizeof(void*)*5, x_241); +lean_ctor_set_uint32(x_247, sizeof(void*)*5 + 4, x_244); lean_inc(x_5); -x_161 = lean_apply_8(x_2, x_3, x_160, x_5, x_6, x_7, x_8, x_9, x_152); -if (lean_obj_tag(x_161) == 0) +x_248 = lean_apply_8(x_2, x_3, x_247, x_5, x_6, x_7, x_8, x_9, x_238); +if (lean_obj_tag(x_248) == 0) { -lean_object* x_162; lean_object* x_163; -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_161, 1); -lean_inc(x_163); -lean_dec(x_161); -x_15 = x_162; -x_16 = x_163; -goto block_35; +lean_object* x_249; lean_object* x_250; +x_249 = lean_ctor_get(x_248, 0); +lean_inc(x_249); +x_250 = lean_ctor_get(x_248, 1); +lean_inc(x_250); +lean_dec(x_248); +x_117 = x_249; +x_118 = x_250; +goto block_138; } else { -lean_object* x_164; lean_object* x_165; -x_164 = lean_ctor_get(x_161, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_161, 1); -lean_inc(x_165); -lean_dec(x_161); -x_36 = x_164; -x_37 = x_165; -goto block_56; +lean_object* x_251; lean_object* x_252; +x_251 = lean_ctor_get(x_248, 0); +lean_inc(x_251); +x_252 = lean_ctor_get(x_248, 1); +lean_inc(x_252); +lean_dec(x_248); +x_139 = x_251; +x_140 = x_252; +goto block_160; } } } else { -lean_object* x_166; lean_object* x_167; +lean_object* x_253; lean_object* x_254; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8641,169 +8955,175 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_166 = lean_ctor_get(x_142, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_142, 1); -lean_inc(x_167); -lean_dec(x_142); -x_36 = x_166; -x_37 = x_167; -goto block_56; -} +x_253 = lean_ctor_get(x_228, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_228, 1); +lean_inc(x_254); +lean_dec(x_228); +x_139 = x_253; +x_140 = x_254; +goto block_160; } } -block_35: +block_138: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = lean_st_ref_take(x_5, 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_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; +x_119 = lean_st_ref_take(x_5, x_118); +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_119, 1); +lean_inc(x_121); +lean_dec(x_119); +x_122 = !lean_is_exclusive(x_120); +if (x_122 == 0) { -lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_21 = lean_ctor_get(x_18, 0); -lean_dec(x_21); -lean_ctor_set(x_18, 0, x_14); -x_22 = lean_st_ref_set(x_5, x_18, x_19); +lean_object* x_123; lean_object* x_124; uint8_t x_125; +x_123 = lean_ctor_get(x_120, 0); +lean_dec(x_123); +lean_ctor_set(x_120, 0, x_116); +x_124 = lean_st_ref_set(x_5, x_120, x_121); lean_dec(x_5); -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_125 = !lean_is_exclusive(x_124); +if (x_125 == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_22, 0); -lean_dec(x_24); -lean_ctor_set(x_22, 0, x_15); -return x_22; +lean_object* x_126; +x_126 = lean_ctor_get(x_124, 0); +lean_dec(x_126); +lean_ctor_set(x_124, 0, x_117); +return x_124; } else { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_22, 1); -lean_inc(x_25); -lean_dec(x_22); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_15); -lean_ctor_set(x_26, 1, x_25); -return x_26; +lean_object* x_127; lean_object* x_128; +x_127 = lean_ctor_get(x_124, 1); +lean_inc(x_127); +lean_dec(x_124); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_117); +lean_ctor_set(x_128, 1, x_127); +return x_128; } } 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; -x_27 = lean_ctor_get(x_18, 1); -x_28 = lean_ctor_get(x_18, 2); -x_29 = lean_ctor_get(x_18, 3); -lean_inc(x_29); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_18); -x_30 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_30, 0, x_14); -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_5, x_30, x_19); +lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_129 = lean_ctor_get(x_120, 1); +x_130 = lean_ctor_get(x_120, 2); +x_131 = lean_ctor_get(x_120, 3); +x_132 = lean_ctor_get(x_120, 4); +lean_inc(x_132); +lean_inc(x_131); +lean_inc(x_130); +lean_inc(x_129); +lean_dec(x_120); +x_133 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_133, 0, x_116); +lean_ctor_set(x_133, 1, x_129); +lean_ctor_set(x_133, 2, x_130); +lean_ctor_set(x_133, 3, x_131); +lean_ctor_set(x_133, 4, x_132); +x_134 = lean_st_ref_set(x_5, x_133, x_121); lean_dec(x_5); -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; +x_135 = lean_ctor_get(x_134, 1); +lean_inc(x_135); +if (lean_is_exclusive(x_134)) { + lean_ctor_release(x_134, 0); + lean_ctor_release(x_134, 1); + x_136 = x_134; } else { - lean_dec_ref(x_31); - x_33 = lean_box(0); + lean_dec_ref(x_134); + x_136 = lean_box(0); } -if (lean_is_scalar(x_33)) { - x_34 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_136)) { + x_137 = lean_alloc_ctor(0, 2, 0); } else { - x_34 = x_33; + x_137 = x_136; } -lean_ctor_set(x_34, 0, x_15); -lean_ctor_set(x_34, 1, x_32); -return x_34; +lean_ctor_set(x_137, 0, x_117); +lean_ctor_set(x_137, 1, x_135); +return x_137; } } -block_56: +block_160: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_38 = lean_st_ref_take(x_5, x_37); -x_39 = lean_ctor_get(x_38, 0); -lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); -x_41 = !lean_is_exclusive(x_39); -if (x_41 == 0) +lean_object* x_141; lean_object* x_142; lean_object* x_143; uint8_t x_144; +x_141 = lean_st_ref_take(x_5, x_140); +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +x_144 = !lean_is_exclusive(x_142); +if (x_144 == 0) { -lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_42 = lean_ctor_get(x_39, 0); -lean_dec(x_42); -lean_ctor_set(x_39, 0, x_14); -x_43 = lean_st_ref_set(x_5, x_39, x_40); +lean_object* x_145; lean_object* x_146; uint8_t x_147; +x_145 = lean_ctor_get(x_142, 0); +lean_dec(x_145); +lean_ctor_set(x_142, 0, x_116); +x_146 = lean_st_ref_set(x_5, x_142, x_143); lean_dec(x_5); -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) +x_147 = !lean_is_exclusive(x_146); +if (x_147 == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_43, 0); -lean_dec(x_45); -lean_ctor_set_tag(x_43, 1); -lean_ctor_set(x_43, 0, x_36); -return x_43; +lean_object* x_148; +x_148 = lean_ctor_get(x_146, 0); +lean_dec(x_148); +lean_ctor_set_tag(x_146, 1); +lean_ctor_set(x_146, 0, x_139); +return x_146; } else { -lean_object* x_46; lean_object* x_47; -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); -lean_dec(x_43); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_36); -lean_ctor_set(x_47, 1, x_46); -return x_47; +lean_object* x_149; lean_object* x_150; +x_149 = lean_ctor_get(x_146, 1); +lean_inc(x_149); +lean_dec(x_146); +x_150 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_150, 0, x_139); +lean_ctor_set(x_150, 1, x_149); +return x_150; } } else { -lean_object* x_48; lean_object* x_49; 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 = lean_ctor_get(x_39, 1); -x_49 = lean_ctor_get(x_39, 2); -x_50 = lean_ctor_get(x_39, 3); -lean_inc(x_50); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_39); -x_51 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_51, 0, x_14); -lean_ctor_set(x_51, 1, x_48); -lean_ctor_set(x_51, 2, x_49); -lean_ctor_set(x_51, 3, x_50); -x_52 = lean_st_ref_set(x_5, x_51, x_40); -lean_dec(x_5); -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; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; +x_151 = lean_ctor_get(x_142, 1); +x_152 = lean_ctor_get(x_142, 2); +x_153 = lean_ctor_get(x_142, 3); +x_154 = lean_ctor_get(x_142, 4); +lean_inc(x_154); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_dec(x_142); +x_155 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_155, 0, x_116); +lean_ctor_set(x_155, 1, x_151); +lean_ctor_set(x_155, 2, x_152); +lean_ctor_set(x_155, 3, x_153); +lean_ctor_set(x_155, 4, x_154); +x_156 = lean_st_ref_set(x_5, x_155, x_143); +lean_dec(x_5); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_158 = x_156; } else { - lean_dec_ref(x_52); - x_54 = lean_box(0); + lean_dec_ref(x_156); + x_158 = lean_box(0); } -if (lean_is_scalar(x_54)) { - x_55 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_158)) { + x_159 = lean_alloc_ctor(1, 2, 0); } else { - x_55 = x_54; - lean_ctor_set_tag(x_55, 1); + x_159 = x_158; + lean_ctor_set_tag(x_159, 1); +} +lean_ctor_set(x_159, 0, x_139); +lean_ctor_set(x_159, 1, x_157); +return x_159; } -lean_ctor_set(x_55, 0, x_36); -lean_ctor_set(x_55, 1, x_53); -return x_55; } } } @@ -12349,29 +12669,32 @@ return x_14; } else { -lean_object* x_15; uint32_t x_16; lean_object* x_17; lean_object* x_18; uint32_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_object* x_15; uint32_t x_16; lean_object* x_17; lean_object* x_18; uint32_t x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_15 = lean_ctor_get(x_3, 0); -x_16 = lean_ctor_get_uint32(x_3, sizeof(void*)*4); +x_16 = lean_ctor_get_uint32(x_3, sizeof(void*)*5); x_17 = lean_ctor_get(x_3, 1); x_18 = lean_ctor_get(x_3, 2); -x_19 = lean_ctor_get_uint32(x_3, sizeof(void*)*4 + 4); +x_19 = lean_ctor_get_uint32(x_3, sizeof(void*)*5 + 4); +x_20 = lean_ctor_get(x_3, 4); +lean_inc(x_20); lean_inc(x_18); lean_inc(x_17); lean_inc(x_15); lean_dec(x_3); lean_inc(x_1); -x_20 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_20, 0, x_1); -x_21 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_17); -lean_ctor_set(x_21, 2, x_18); -lean_ctor_set(x_21, 3, x_20); -lean_ctor_set_uint32(x_21, sizeof(void*)*4, x_16); -lean_ctor_set_uint32(x_21, sizeof(void*)*4 + 4, x_19); -x_22 = l_Lean_Meta_Simp_simpLambda___closed__1; -x_23 = l_Lean_Meta_Simp_lambdaTelescopeDSimp___rarg(x_1, x_22, x_2, x_21, x_4, x_5, x_6, x_7, x_8, x_9); -return x_23; +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_1); +x_22 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_22, 0, x_15); +lean_ctor_set(x_22, 1, x_17); +lean_ctor_set(x_22, 2, x_18); +lean_ctor_set(x_22, 3, x_21); +lean_ctor_set(x_22, 4, x_20); +lean_ctor_set_uint32(x_22, sizeof(void*)*5, x_16); +lean_ctor_set_uint32(x_22, sizeof(void*)*5 + 4, x_19); +x_23 = l_Lean_Meta_Simp_simpLambda___closed__1; +x_24 = l_Lean_Meta_Simp_lambdaTelescopeDSimp___rarg(x_1, x_23, x_2, x_22, x_4, x_5, x_6, x_7, x_8, x_9); +return x_24; } } } @@ -16287,65 +16610,68 @@ return x_29; } else { -lean_object* x_30; uint32_t x_31; lean_object* x_32; lean_object* x_33; uint32_t 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_30; uint32_t x_31; lean_object* x_32; lean_object* x_33; uint32_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; x_30 = lean_ctor_get(x_3, 0); -x_31 = lean_ctor_get_uint32(x_3, sizeof(void*)*4); +x_31 = lean_ctor_get_uint32(x_3, sizeof(void*)*5); x_32 = lean_ctor_get(x_3, 1); x_33 = lean_ctor_get(x_3, 2); -x_34 = lean_ctor_get_uint32(x_3, sizeof(void*)*4 + 4); +x_34 = lean_ctor_get_uint32(x_3, sizeof(void*)*5 + 4); +x_35 = lean_ctor_get(x_3, 4); +lean_inc(x_35); lean_inc(x_33); lean_inc(x_32); lean_inc(x_30); lean_dec(x_3); lean_inc(x_1); -x_35 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_35, 0, x_1); -x_36 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_36, 0, x_30); -lean_ctor_set(x_36, 1, x_32); -lean_ctor_set(x_36, 2, x_33); -lean_ctor_set(x_36, 3, x_35); -lean_ctor_set_uint32(x_36, sizeof(void*)*4, x_31); -lean_ctor_set_uint32(x_36, sizeof(void*)*4 + 4, x_34); -x_37 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___spec__1(x_10, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_9); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_unbox(x_38); -lean_dec(x_38); -if (x_39 == 0) +x_36 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_1); +x_37 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_37, 0, x_30); +lean_ctor_set(x_37, 1, x_32); +lean_ctor_set(x_37, 2, x_33); +lean_ctor_set(x_37, 3, x_36); +lean_ctor_set(x_37, 4, x_35); +lean_ctor_set_uint32(x_37, sizeof(void*)*5, x_31); +lean_ctor_set_uint32(x_37, sizeof(void*)*5 + 4, x_34); +x_38 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___spec__1(x_10, x_2, x_37, x_4, x_5, x_6, x_7, x_8, x_9); +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_unbox(x_39); +lean_dec(x_39); +if (x_40 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_37, 1); -lean_inc(x_40); -lean_dec(x_37); -x_41 = lean_box(0); -x_42 = l_Lean_Meta_Simp_simpForall___lambda__6(x_1, x_41, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_40); -return x_42; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_38, 1); +lean_inc(x_41); +lean_dec(x_38); +x_42 = lean_box(0); +x_43 = l_Lean_Meta_Simp_simpForall___lambda__6(x_1, x_42, x_2, x_37, x_4, x_5, x_6, x_7, x_8, x_41); +return x_43; } 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; -x_43 = lean_ctor_get(x_37, 1); -lean_inc(x_43); -lean_dec(x_37); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_44 = lean_ctor_get(x_38, 1); +lean_inc(x_44); +lean_dec(x_38); lean_inc(x_1); -x_44 = l_Lean_MessageData_ofExpr(x_1); -x_45 = l_Lean_Meta_Simp_simpForall___closed__2; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -x_47 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_addTrace___at_Lean_Meta_Simp_congrArgs___spec__2(x_10, x_48, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_43); -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); +x_45 = l_Lean_MessageData_ofExpr(x_1); +x_46 = l_Lean_Meta_Simp_simpForall___closed__2; +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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +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_Meta_Simp_congrArgs___spec__2(x_10, x_49, x_2, x_37, x_4, x_5, x_6, x_7, x_8, x_44); +x_51 = lean_ctor_get(x_50, 0); lean_inc(x_51); -lean_dec(x_49); -x_52 = l_Lean_Meta_Simp_simpForall___lambda__6(x_1, x_50, x_2, x_36, x_4, x_5, x_6, x_7, x_8, x_51); -return x_52; +x_52 = lean_ctor_get(x_50, 1); +lean_inc(x_52); +lean_dec(x_50); +x_53 = l_Lean_Meta_Simp_simpForall___lambda__6(x_1, x_51, x_2, x_37, x_4, x_5, x_6, x_7, x_8, x_52); +return x_53; } } } @@ -18251,7 +18577,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(391u); +x_3 = lean_unsigned_to_nat(390u); 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); @@ -23020,90 +23346,93 @@ return x_22; } else { -lean_object* x_23; uint32_t x_24; lean_object* x_25; lean_object* x_26; uint32_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +lean_object* x_23; uint32_t x_24; lean_object* x_25; lean_object* x_26; uint32_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; x_23 = lean_ctor_get(x_3, 0); -x_24 = lean_ctor_get_uint32(x_3, sizeof(void*)*4); +x_24 = lean_ctor_get_uint32(x_3, sizeof(void*)*5); x_25 = lean_ctor_get(x_3, 1); x_26 = lean_ctor_get(x_3, 2); -x_27 = lean_ctor_get_uint32(x_3, sizeof(void*)*4 + 4); +x_27 = lean_ctor_get_uint32(x_3, sizeof(void*)*5 + 4); +x_28 = lean_ctor_get(x_3, 4); +lean_inc(x_28); lean_inc(x_26); lean_inc(x_25); lean_inc(x_23); lean_dec(x_3); lean_inc(x_1); -x_28 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_28, 0, x_1); -x_29 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_29, 0, x_23); -lean_ctor_set(x_29, 1, x_25); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_28); -lean_ctor_set_uint32(x_29, sizeof(void*)*4, x_24); -lean_ctor_set_uint32(x_29, sizeof(void*)*4 + 4, x_27); -x_30 = l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_congrDefault___spec__1(x_1, x_16, x_18, x_2, x_29, x_4, x_5, x_6, x_7, x_8, x_12); -return x_30; +x_29 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_29, 0, x_1); +x_30 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_30, 0, x_23); +lean_ctor_set(x_30, 1, x_25); +lean_ctor_set(x_30, 2, x_26); +lean_ctor_set(x_30, 3, x_29); +lean_ctor_set(x_30, 4, x_28); +lean_ctor_set_uint32(x_30, sizeof(void*)*5, x_24); +lean_ctor_set_uint32(x_30, sizeof(void*)*5 + 4, x_27); +x_31 = l_Lean_Expr_withAppAux___at_Lean_Meta_Simp_congrDefault___spec__1(x_1, x_16, x_18, x_2, x_30, x_4, x_5, x_6, x_7, x_8, x_12); +return x_31; } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_dec(x_1); -x_31 = lean_ctor_get(x_10, 1); -lean_inc(x_31); -lean_dec(x_10); -x_32 = lean_ctor_get(x_11, 0); +x_32 = lean_ctor_get(x_10, 1); lean_inc(x_32); -lean_dec(x_11); -x_33 = lean_ctor_get(x_32, 0); +lean_dec(x_10); +x_33 = lean_ctor_get(x_11, 0); lean_inc(x_33); +lean_dec(x_11); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_34 = l_Lean_Meta_Simp_visitFn(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_34) == 0) +x_35 = l_Lean_Meta_Simp_visitFn(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_32); +if (lean_obj_tag(x_35) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_Meta_Simp_Result_mkEqTrans(x_32, x_35, x_5, x_6, x_7, x_8, x_36); -return x_37; +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_Lean_Meta_Simp_Result_mkEqTrans(x_33, x_36, x_5, x_6, x_7, x_8, x_37); +return x_38; } else { -uint8_t x_38; -lean_dec(x_32); +uint8_t x_39; +lean_dec(x_33); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_38 = !lean_is_exclusive(x_34); -if (x_38 == 0) +x_39 = !lean_is_exclusive(x_35); +if (x_39 == 0) { -return x_34; +return x_35; } else { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_39 = lean_ctor_get(x_34, 0); -x_40 = lean_ctor_get(x_34, 1); +lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_35, 0); +x_41 = lean_ctor_get(x_35, 1); +lean_inc(x_41); lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_34); -x_41 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -return x_41; +lean_dec(x_35); +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_42; +uint8_t x_43; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -23112,23 +23441,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_42 = !lean_is_exclusive(x_10); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_10); +if (x_43 == 0) { return x_10; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_10, 0); -x_44 = lean_ctor_get(x_10, 1); +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_10, 0); +x_45 = lean_ctor_get(x_10, 1); +lean_inc(x_45); lean_inc(x_44); -lean_inc(x_43); lean_dec(x_10); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -return x_45; +x_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; } } } @@ -25722,20 +26051,23 @@ return x_27; } else { -lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; lean_object* x_35; +lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; size_t x_34; size_t x_35; lean_object* x_36; x_28 = lean_ctor_get(x_21, 1); lean_inc(x_28); lean_dec(x_21); -x_29 = 0; -x_30 = lean_box(x_29); +x_29 = lean_ctor_get(x_3, 2); +lean_inc(x_29); +lean_dec(x_3); +x_30 = 0; +x_31 = lean_box(x_30); lean_inc(x_2); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_2); -lean_ctor_set(x_31, 1, x_30); -x_32 = lean_array_get_size(x_3); -x_33 = lean_usize_of_nat(x_32); -lean_dec(x_32); -x_34 = 0; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_2); +lean_ctor_set(x_32, 1, x_31); +x_33 = lean_array_get_size(x_29); +x_34 = lean_usize_of_nat(x_33); +lean_dec(x_33); +x_35 = 0; lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); @@ -25745,34 +26077,34 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_2); lean_inc(x_4); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__1(x_4, x_5, x_2, x_3, x_33, x_34, x_31, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_28); -lean_dec(x_3); -if (lean_obj_tag(x_35) == 0) +x_36 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__1(x_4, x_5, x_2, x_29, x_34, x_35, x_32, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_28); +lean_dec(x_29); +if (lean_obj_tag(x_36) == 0) { -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); +lean_object* x_37; lean_object* x_38; x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -if (lean_obj_tag(x_37) == 0) -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; -x_38 = lean_ctor_get(x_35, 1); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -lean_dec(x_35); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; lean_object* x_43; x_39 = lean_ctor_get(x_36, 1); lean_inc(x_39); lean_dec(x_36); -x_40 = lean_box(0); -x_41 = lean_unbox(x_39); -lean_dec(x_39); -x_42 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5(x_41, x_2, x_4, x_6, x_5, x_7, x_8, x_11, x_9, x_40, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_38); -return x_42; +x_40 = lean_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = lean_box(0); +x_42 = lean_unbox(x_40); +lean_dec(x_40); +x_43 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5(x_42, x_2, x_4, x_6, x_5, x_7, x_8, x_11, x_9, x_41, x_13, x_14, x_15, x_16, x_17, x_18, x_19, x_39); +return x_43; } else { -uint8_t x_43; -lean_dec(x_36); +uint8_t x_44; +lean_dec(x_37); lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -25787,37 +26119,37 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_43 = !lean_is_exclusive(x_35); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_36); +if (x_44 == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_35, 0); -lean_dec(x_44); -x_45 = lean_ctor_get(x_37, 0); -lean_inc(x_45); -lean_dec(x_37); -lean_ctor_set(x_35, 0, x_45); -return x_35; +lean_object* x_45; lean_object* x_46; +x_45 = lean_ctor_get(x_36, 0); +lean_dec(x_45); +x_46 = lean_ctor_get(x_38, 0); +lean_inc(x_46); +lean_dec(x_38); +lean_ctor_set(x_36, 0, x_46); +return x_36; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_35, 1); -lean_inc(x_46); -lean_dec(x_35); -x_47 = lean_ctor_get(x_37, 0); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_36, 1); lean_inc(x_47); -lean_dec(x_37); -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_dec(x_36); +x_48 = lean_ctor_get(x_38, 0); +lean_inc(x_48); +lean_dec(x_38); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +return x_49; } } } else { -uint8_t x_49; +uint8_t x_50; lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -25832,30 +26164,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_49 = !lean_is_exclusive(x_35); -if (x_49 == 0) +x_50 = !lean_is_exclusive(x_36); +if (x_50 == 0) { -return x_35; +return x_36; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_35, 0); -x_51 = lean_ctor_get(x_35, 1); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_36, 0); +x_52 = lean_ctor_get(x_36, 1); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_35); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_dec(x_36); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_51); +lean_ctor_set(x_53, 1, x_52); +return x_53; } } } } else { -uint8_t x_53; +uint8_t x_54; lean_dec(x_19); lean_dec(x_18); lean_dec(x_17); @@ -25871,23 +26203,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_53 = !lean_is_exclusive(x_21); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_21); +if (x_54 == 0) { return x_21; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_21, 0); -x_55 = lean_ctor_get(x_21, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_21, 0); +x_56 = lean_ctor_get(x_21, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); lean_dec(x_21); -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; +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; } } } @@ -25946,84 +26278,82 @@ return x_43; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___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_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___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) { _start: { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -lean_dec(x_3); -x_12 = lean_ctor_get(x_1, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_1, 2); -lean_inc(x_13); -lean_dec(x_1); +lean_object* x_13; +lean_dec(x_4); +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_12); -x_14 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_12, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_14) == 0) +lean_inc(x_1); +x_13 = l_Lean_Meta_mkConstWithFreshMVarLevels(x_1, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(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_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); +lean_dec(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_15); -x_17 = lean_infer_type(x_15, x_7, x_8, x_9, x_10, x_16); -if (lean_obj_tag(x_17) == 0) +lean_inc(x_14); +x_16 = lean_infer_type(x_14, x_8, x_9, x_10, x_11, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; lean_object* x_23; -x_18 = lean_ctor_get(x_17, 0); +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t 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); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); -x_21 = 1; -x_22 = 0; +lean_dec(x_16); +x_19 = lean_box(0); +x_20 = 1; +x_21 = 0; +lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -x_23 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(x_18, x_21, x_20, x_22, x_7, x_8, x_9, x_10, x_19); -if (lean_obj_tag(x_23) == 0) +x_22 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(x_17, x_20, x_19, x_21, x_8, x_9, x_10, x_11, x_18); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_24; lean_object* x_25; uint8_t x_26; -x_24 = lean_ctor_get(x_23, 0); +lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_23, 1); lean_inc(x_24); -x_25 = lean_ctor_get(x_24, 1); -lean_inc(x_25); -x_26 = !lean_is_exclusive(x_23); -if (x_26 == 0) +x_25 = !lean_is_exclusive(x_22); +if (x_25 == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_27 = lean_ctor_get(x_23, 1); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +x_26 = lean_ctor_get(x_22, 1); +x_27 = lean_ctor_get(x_22, 0); +lean_dec(x_27); x_28 = lean_ctor_get(x_23, 0); -lean_dec(x_28); +lean_inc(x_28); +lean_dec(x_23); x_29 = lean_ctor_get(x_24, 0); lean_inc(x_29); -lean_dec(x_24); -x_30 = lean_ctor_get(x_25, 0); +x_30 = lean_ctor_get(x_24, 1); lean_inc(x_30); -x_31 = lean_ctor_get(x_25, 1); +lean_dec(x_24); +x_31 = lean_ctor_get(x_2, 2); lean_inc(x_31); -lean_dec(x_25); -x_32 = lean_array_get_size(x_13); +x_32 = lean_array_get_size(x_31); x_33 = lean_unsigned_to_nat(0u); x_34 = lean_nat_dec_lt(x_33, x_32); if (x_34 == 0) { lean_object* x_35; lean_object* x_36; lean_dec(x_32); -lean_free_object(x_23); +lean_dec(x_31); +lean_free_object(x_22); x_35 = lean_box(0); -x_36 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_31, x_20, x_13, x_12, x_29, x_30, x_15, x_2, x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); +x_36 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_30, x_19, x_2, x_1, x_28, x_29, x_14, x_3, x_35, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); return x_36; } else @@ -26032,247 +26362,208 @@ size_t x_37; size_t x_38; uint8_t x_39; x_37 = 0; x_38 = lean_usize_of_nat(x_32); lean_dec(x_32); -x_39 = l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__2(x_29, x_13, x_37, x_38); +x_39 = l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__2(x_28, x_31, x_37, x_38); +lean_dec(x_31); if (x_39 == 0) { lean_object* x_40; lean_object* x_41; -lean_free_object(x_23); +lean_free_object(x_22); x_40 = lean_box(0); -x_41 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_31, x_20, x_13, x_12, x_29, x_30, x_15, x_2, x_40, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); +x_41 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_30, x_19, x_2, x_1, x_28, x_29, x_14, x_3, x_40, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); return x_41; } else { -lean_dec(x_31); lean_dec(x_30); lean_dec(x_29); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_28); +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); -lean_ctor_set(x_23, 0, x_20); -return x_23; +lean_dec(x_1); +lean_ctor_set(x_22, 0, x_19); +return x_22; } } } 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; uint8_t x_48; -x_42 = lean_ctor_get(x_23, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; +x_42 = lean_ctor_get(x_22, 1); lean_inc(x_42); -lean_dec(x_23); -x_43 = lean_ctor_get(x_24, 0); +lean_dec(x_22); +x_43 = lean_ctor_get(x_23, 0); lean_inc(x_43); -lean_dec(x_24); -x_44 = lean_ctor_get(x_25, 0); +lean_dec(x_23); +x_44 = lean_ctor_get(x_24, 0); lean_inc(x_44); -x_45 = lean_ctor_get(x_25, 1); +x_45 = lean_ctor_get(x_24, 1); lean_inc(x_45); -lean_dec(x_25); -x_46 = lean_array_get_size(x_13); -x_47 = lean_unsigned_to_nat(0u); -x_48 = lean_nat_dec_lt(x_47, x_46); -if (x_48 == 0) +lean_dec(x_24); +x_46 = lean_ctor_get(x_2, 2); +lean_inc(x_46); +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) { -lean_object* x_49; lean_object* x_50; +lean_object* x_50; lean_object* x_51; +lean_dec(x_47); lean_dec(x_46); -x_49 = lean_box(0); -x_50 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_45, x_20, x_13, x_12, x_43, x_44, x_15, x_2, x_49, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); -return x_50; +x_50 = lean_box(0); +x_51 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_45, x_19, x_2, x_1, x_43, x_44, x_14, x_3, x_50, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +return x_51; } else { -size_t x_51; size_t x_52; uint8_t x_53; -x_51 = 0; -x_52 = lean_usize_of_nat(x_46); +size_t x_52; size_t x_53; uint8_t x_54; +x_52 = 0; +x_53 = lean_usize_of_nat(x_47); +lean_dec(x_47); +x_54 = l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__2(x_43, x_46, x_52, x_53); lean_dec(x_46); -x_53 = l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_trySimpCongrTheorem_x3f___spec__2(x_43, x_13, x_51, x_52); -if (x_53 == 0) +if (x_54 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_box(0); -x_55 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_45, x_20, x_13, x_12, x_43, x_44, x_15, x_2, x_54, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_42); -return x_55; +lean_object* x_55; lean_object* x_56; +x_55 = lean_box(0); +x_56 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_45, x_19, x_2, x_1, x_43, x_44, x_14, x_3, x_55, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_42); +return x_56; } else { -lean_object* x_56; +lean_object* x_57; lean_dec(x_45); lean_dec(x_44); lean_dec(x_43); -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); +lean_dec(x_14); +lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_4); +lean_dec(x_3); lean_dec(x_2); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_20); -lean_ctor_set(x_56, 1, x_42); -return x_56; +lean_dec(x_1); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_19); +lean_ctor_set(x_57, 1, x_42); +return x_57; } } } } else { -uint8_t x_57; -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_58; +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_57 = !lean_is_exclusive(x_23); -if (x_57 == 0) +lean_dec(x_1); +x_58 = !lean_is_exclusive(x_22); +if (x_58 == 0) { -return x_23; +return x_22; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_23, 0); -x_59 = lean_ctor_get(x_23, 1); +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_inc(x_58); -lean_dec(x_23); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_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; } } } else { -uint8_t x_61; -lean_dec(x_15); -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_62; +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_61 = !lean_is_exclusive(x_17); -if (x_61 == 0) +lean_dec(x_1); +x_62 = !lean_is_exclusive(x_16); +if (x_62 == 0) { -return x_17; +return x_16; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_17, 0); -x_63 = lean_ctor_get(x_17, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_16, 0); +x_64 = lean_ctor_get(x_16, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_17); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_63); -return x_64; +lean_dec(x_16); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } } else { -uint8_t x_65; -lean_dec(x_13); -lean_dec(x_12); +uint8_t x_66; +lean_dec(x_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_65 = !lean_is_exclusive(x_14); -if (x_65 == 0) +lean_dec(x_1); +x_66 = !lean_is_exclusive(x_13); +if (x_66 == 0) { -return x_14; +return x_13; } else { -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_object* x_67; lean_object* x_68; lean_object* x_69; +x_67 = lean_ctor_get(x_13, 0); +x_68 = lean_ctor_get(x_13, 1); +lean_inc(x_68); lean_inc(x_67); -lean_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_dec(x_13); +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_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -if (x_4 == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_3); -x_13 = lean_box(0); -x_14 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(x_1, x_2, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_15 = lean_ctor_get(x_1, 0); -lean_inc(x_15); -x_16 = l_Lean_MessageData_ofName(x_15); -x_17 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_16); -x_19 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__9; -x_20 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -lean_inc(x_2); -x_21 = l_Lean_MessageData_ofExpr(x_2); -x_22 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_22, 0, x_20); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_17); -x_24 = l_Lean_addTrace___at_Lean_Meta_Simp_congrArgs___spec__2(x_3, x_23, 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); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(x_1, x_2, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); -return x_27; -} -} -} -static lean_object* _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___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; @@ -26285,29 +26576,78 @@ x_6 = l_Lean_Name_mkStr5(x_1, x_2, x_3, x_4, x_5); return x_6; } } -static lean_object* _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___spec__1___boxed), 9, 1); -lean_closure_set(x_2, 0, x_1); -return x_2; +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +lean_dec(x_4); +x_13 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___closed__1; +x_14 = l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___spec__1(x_13, 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_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = lean_box(0); +x_19 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(x_1, x_2, x_3, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_19; +} +else +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_20 = lean_ctor_get(x_14, 1); +lean_inc(x_20); +lean_dec(x_14); +lean_inc(x_1); +x_21 = l_Lean_MessageData_ofName(x_1); +x_22 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__9; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +lean_inc(x_3); +x_26 = l_Lean_MessageData_ofExpr(x_3); +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_28, 0, x_27); +lean_ctor_set(x_28, 1, x_22); +x_29 = l_Lean_addTrace___at_Lean_Meta_Simp_congrArgs___spec__2(x_13, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_20); +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_Meta_Simp_trySimpCongrTheorem_x3f___lambda__8(x_1, x_2, x_3, x_30, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_31); +return x_32; +} } } LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; -x_11 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1; -x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___boxed), 12, 3); -lean_closure_set(x_12, 0, x_1); -lean_closure_set(x_12, 1, x_2); -lean_closure_set(x_12, 2, x_11); -x_13 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2; +x_11 = lean_ctor_get(x_1, 0); +lean_inc(x_11); +lean_inc(x_11); +x_12 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_recordCongrTheorem___boxed), 9, 1); +lean_closure_set(x_12, 0, x_11); +x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9), 12, 3); +lean_closure_set(x_13, 0, x_11); +lean_closure_set(x_13, 1, x_1); +lean_closure_set(x_13, 2, x_2); x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Meta_Simp_discharge_x3f_x27___spec__5___rarg), 10, 2); -lean_closure_set(x_14, 0, x_13); -lean_closure_set(x_14, 1, x_12); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); x_15 = 0; x_16 = l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_Simp_tryTheoremWithExtraArgs_x3f___spec__1___rarg(x_14, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_16; @@ -26492,16 +26832,6 @@ x_18 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__7(x_1, x_2, x_3, x_4, return x_18; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: -{ -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_4); -lean_dec(x_4); -x_14 = l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; -} -} LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Meta_Simp_congr___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: { @@ -28224,42 +28554,45 @@ return x_26; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; 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_27 = lean_ctor_get(x_17, 0); x_28 = lean_ctor_get(x_17, 1); x_29 = lean_ctor_get(x_17, 2); x_30 = lean_ctor_get(x_17, 3); +x_31 = lean_ctor_get(x_17, 4); +lean_inc(x_31); lean_inc(x_30); lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_dec(x_17); lean_inc(x_3); -x_31 = l_Lean_SMap_insert___at_Lean_Meta_Simp_cacheResult___spec__1(x_27, x_1, x_3); -x_32 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_28); -lean_ctor_set(x_32, 2, x_29); -lean_ctor_set(x_32, 3, x_30); -x_33 = lean_st_ref_set(x_6, x_32, x_18); -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; +x_32 = l_Lean_SMap_insert___at_Lean_Meta_Simp_cacheResult___spec__1(x_27, x_1, x_3); +x_33 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_28); +lean_ctor_set(x_33, 2, x_29); +lean_ctor_set(x_33, 3, x_30); +lean_ctor_set(x_33, 4, x_31); +x_34 = lean_st_ref_set(x_6, x_33, x_18); +x_35 = lean_ctor_get(x_34, 1); +lean_inc(x_35); +if (lean_is_exclusive(x_34)) { + lean_ctor_release(x_34, 0); + lean_ctor_release(x_34, 1); + x_36 = x_34; } else { - lean_dec_ref(x_33); - x_35 = lean_box(0); + lean_dec_ref(x_34); + x_36 = lean_box(0); } -if (lean_is_scalar(x_35)) { - x_36 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_36)) { + x_37 = lean_alloc_ctor(0, 2, 0); } else { - x_36 = x_35; + x_37 = x_36; } -lean_ctor_set(x_36, 0, x_3); -lean_ctor_set(x_36, 1, x_34); -return x_36; +lean_ctor_set(x_37, 0, x_3); +lean_ctor_set(x_37, 1, x_35); +return x_37; } } } @@ -29515,27 +29848,30 @@ return x_93; } else { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +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_94 = lean_ctor_get(x_51, 0); x_95 = lean_ctor_get(x_51, 1); x_96 = lean_ctor_get(x_51, 2); x_97 = lean_ctor_get(x_51, 3); +x_98 = lean_ctor_get(x_51, 4); +lean_inc(x_98); lean_inc(x_97); lean_inc(x_96); lean_inc(x_95); lean_inc(x_94); lean_dec(x_51); -x_98 = lean_nat_add(x_97, x_36); +x_99 = lean_nat_add(x_97, x_36); lean_dec(x_97); -x_99 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_99, 0, x_94); -lean_ctor_set(x_99, 1, x_95); -lean_ctor_set(x_99, 2, x_96); -lean_ctor_set(x_99, 3, x_98); -x_100 = lean_st_ref_set(x_4, x_99, x_52); -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -lean_dec(x_100); +x_100 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_100, 0, x_94); +lean_ctor_set(x_100, 1, x_95); +lean_ctor_set(x_100, 2, x_96); +lean_ctor_set(x_100, 3, x_99); +lean_ctor_set(x_100, 4, x_98); +x_101 = lean_st_ref_set(x_4, x_100, x_52); +x_102 = lean_ctor_get(x_101, 1); +lean_inc(x_102); +lean_dec(x_101); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -29544,23 +29880,23 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_102 = l_Lean_Meta_Simp_pre(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_101); -if (lean_obj_tag(x_102) == 0) +x_103 = l_Lean_Meta_Simp_pre(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_102); +if (lean_obj_tag(x_103) == 0) { -lean_object* x_103; -x_103 = lean_ctor_get(x_102, 0); -lean_inc(x_103); -switch (lean_obj_tag(x_103)) { +lean_object* x_104; +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +switch (lean_obj_tag(x_104)) { case 0: { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_102, 1); -lean_inc(x_104); -lean_dec(x_102); -x_105 = lean_ctor_get(x_103, 0); +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_103, 1); lean_inc(x_105); lean_dec(x_103); -x_106 = l_Lean_Meta_Simp_cacheResult(x_1, x_39, x_105, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); +x_106 = lean_ctor_get(x_104, 0); +lean_inc(x_106); +lean_dec(x_104); +x_107 = l_Lean_Meta_Simp_cacheResult(x_1, x_39, x_106, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_105); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -29569,19 +29905,19 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_39); -return x_106; +return x_107; } case 1: { -lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; -x_107 = lean_ctor_get(x_102, 1); -lean_inc(x_107); -lean_dec(x_102); -x_108 = lean_ctor_get(x_103, 0); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_103, 1); lean_inc(x_108); lean_dec(x_103); -x_109 = lean_ctor_get(x_108, 0); +x_109 = lean_ctor_get(x_104, 0); lean_inc(x_109); +lean_dec(x_104); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -29589,29 +29925,29 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_110 = l_Lean_Meta_Simp_simpLoop(x_109, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_107); -if (lean_obj_tag(x_110) == 0) +x_111 = l_Lean_Meta_Simp_simpLoop(x_110, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_108); +if (lean_obj_tag(x_111) == 0) { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_110, 1); +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_111, 0); lean_inc(x_112); -lean_dec(x_110); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_113 = l_Lean_Meta_Simp_Result_mkEqTrans(x_108, x_111, x_5, x_6, x_7, x_8, x_112); -if (lean_obj_tag(x_113) == 0) +x_114 = l_Lean_Meta_Simp_Result_mkEqTrans(x_109, x_112, x_5, x_6, x_7, x_8, x_113); +if (lean_obj_tag(x_114) == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_114, 0); lean_inc(x_115); -lean_dec(x_113); -x_116 = l_Lean_Meta_Simp_cacheResult(x_1, x_39, x_114, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +lean_dec(x_114); +x_117 = l_Lean_Meta_Simp_cacheResult(x_1, x_39, x_115, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_116); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -29620,11 +29956,11 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_39); -return x_116; +return x_117; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_dec(x_39); lean_dec(x_7); lean_dec(x_8); @@ -29634,32 +29970,32 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_117 = lean_ctor_get(x_113, 0); -lean_inc(x_117); -x_118 = lean_ctor_get(x_113, 1); +x_118 = lean_ctor_get(x_114, 0); lean_inc(x_118); -if (lean_is_exclusive(x_113)) { - lean_ctor_release(x_113, 0); - lean_ctor_release(x_113, 1); - x_119 = x_113; +x_119 = lean_ctor_get(x_114, 1); +lean_inc(x_119); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_120 = x_114; } else { - lean_dec_ref(x_113); - x_119 = lean_box(0); + lean_dec_ref(x_114); + x_120 = lean_box(0); } -if (lean_is_scalar(x_119)) { - x_120 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_120)) { + x_121 = lean_alloc_ctor(1, 2, 0); } else { - x_120 = x_119; + x_121 = x_120; } -lean_ctor_set(x_120, 0, x_117); -lean_ctor_set(x_120, 1, x_118); -return x_120; +lean_ctor_set(x_121, 0, x_118); +lean_ctor_set(x_121, 1, x_119); +return x_121; } } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_dec(x_108); +lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_dec(x_109); lean_dec(x_39); lean_dec(x_7); lean_dec(x_8); @@ -29669,70 +30005,70 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_121 = lean_ctor_get(x_110, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_110, 1); +x_122 = lean_ctor_get(x_111, 0); lean_inc(x_122); -if (lean_is_exclusive(x_110)) { - lean_ctor_release(x_110, 0); - lean_ctor_release(x_110, 1); - x_123 = x_110; +x_123 = lean_ctor_get(x_111, 1); +lean_inc(x_123); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_124 = x_111; } else { - lean_dec_ref(x_110); - x_123 = lean_box(0); + lean_dec_ref(x_111); + x_124 = lean_box(0); } -if (lean_is_scalar(x_123)) { - x_124 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_124)) { + x_125 = lean_alloc_ctor(1, 2, 0); } else { - x_124 = x_123; + x_125 = x_124; } -lean_ctor_set(x_124, 0, x_121); -lean_ctor_set(x_124, 1, x_122); -return x_124; +lean_ctor_set(x_125, 0, x_122); +lean_ctor_set(x_125, 1, x_123); +return x_125; } } default: { -lean_object* x_125; -x_125 = lean_ctor_get(x_103, 0); -lean_inc(x_125); -lean_dec(x_103); -if (lean_obj_tag(x_125) == 0) -{ -lean_object* x_126; lean_object* x_127; uint8_t x_128; lean_object* x_129; lean_object* x_130; -x_126 = lean_ctor_get(x_102, 1); +lean_object* x_126; +x_126 = lean_ctor_get(x_104, 0); lean_inc(x_126); -lean_dec(x_102); -x_127 = lean_box(0); -x_128 = 1; +lean_dec(x_104); +if (lean_obj_tag(x_126) == 0) +{ +lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +x_127 = lean_ctor_get(x_103, 1); +lean_inc(x_127); +lean_dec(x_103); +x_128 = lean_box(0); +x_129 = 1; lean_inc(x_1); -x_129 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_129, 0, x_1); -lean_ctor_set(x_129, 1, x_127); -lean_ctor_set_uint8(x_129, sizeof(void*)*2, x_128); -x_130 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_39, x_129, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_126); +x_130 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_130, 0, x_1); +lean_ctor_set(x_130, 1, x_128); +lean_ctor_set_uint8(x_130, sizeof(void*)*2, x_129); +x_131 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_39, x_130, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_127); lean_dec(x_39); -return x_130; +return x_131; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_102, 1); -lean_inc(x_131); -lean_dec(x_102); -x_132 = lean_ctor_get(x_125, 0); +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_103, 1); lean_inc(x_132); -lean_dec(x_125); -x_133 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_39, x_132, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_131); +lean_dec(x_103); +x_133 = lean_ctor_get(x_126, 0); +lean_inc(x_133); +lean_dec(x_126); +x_134 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_39, x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_132); lean_dec(x_39); -return x_133; +return x_134; } } } } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_dec(x_39); lean_dec(x_7); lean_dec(x_8); @@ -29742,32 +30078,32 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_134 = lean_ctor_get(x_102, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_102, 1); +x_135 = lean_ctor_get(x_103, 0); lean_inc(x_135); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_136 = x_102; +x_136 = lean_ctor_get(x_103, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_103)) { + lean_ctor_release(x_103, 0); + lean_ctor_release(x_103, 1); + x_137 = x_103; } else { - lean_dec_ref(x_102); - x_136 = lean_box(0); + lean_dec_ref(x_103); + x_137 = lean_box(0); } -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); } else { - x_137 = x_136; + x_138 = x_137; } -lean_ctor_set(x_137, 0, x_134); -lean_ctor_set(x_137, 1, x_135); -return x_137; +lean_ctor_set(x_138, 0, x_135); +lean_ctor_set(x_138, 1, x_136); +return x_138; } } } else { -uint8_t x_138; +uint8_t x_139; lean_dec(x_39); lean_dec(x_7); lean_dec(x_8); @@ -29777,33 +30113,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_138 = !lean_is_exclusive(x_48); -if (x_138 == 0) +x_139 = !lean_is_exclusive(x_48); +if (x_139 == 0) { return x_48; } else { -lean_object* x_139; lean_object* x_140; lean_object* x_141; -x_139 = lean_ctor_get(x_48, 0); -x_140 = lean_ctor_get(x_48, 1); +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_48, 0); +x_141 = lean_ctor_get(x_48, 1); +lean_inc(x_141); lean_inc(x_140); -lean_inc(x_139); lean_dec(x_48); -x_141 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_141, 0, x_139); -lean_ctor_set(x_141, 1, x_140); -return x_141; +x_142 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +return x_142; } } } else { -lean_object* x_142; lean_object* x_143; +lean_object* x_143; lean_object* x_144; lean_dec(x_39); lean_dec(x_1); -x_142 = l_Lean_Meta_Simp_simpLoop___closed__2; -x_143 = l_Lean_throwError___at_Lean_Meta_Simp_simpLoop___spec__1(x_142, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_43); +x_143 = l_Lean_Meta_Simp_simpLoop___closed__2; +x_144 = l_Lean_throwError___at_Lean_Meta_Simp_simpLoop___spec__1(x_143, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_43); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -29811,191 +30147,195 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_143; +return x_144; } } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; +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; uint8_t x_156; lean_dec(x_7); -x_144 = lean_unsigned_to_nat(1u); -x_145 = lean_nat_add(x_13, x_144); +x_145 = lean_unsigned_to_nat(1u); +x_146 = lean_nat_add(x_13, x_145); lean_dec(x_13); -x_146 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_146, 0, x_10); -lean_ctor_set(x_146, 1, x_11); -lean_ctor_set(x_146, 2, x_12); -lean_ctor_set(x_146, 3, x_145); -lean_ctor_set(x_146, 4, x_14); -lean_ctor_set(x_146, 5, x_15); -lean_ctor_set(x_146, 6, x_16); -lean_ctor_set(x_146, 7, x_17); -lean_ctor_set(x_146, 8, x_18); -lean_ctor_set(x_146, 9, x_19); -lean_ctor_set(x_146, 10, x_20); -lean_ctor_set_uint8(x_146, sizeof(void*)*11, x_21); -lean_ctor_set_uint8(x_146, sizeof(void*)*11 + 1, x_22); -x_147 = l_Lean_Meta_Simp_getConfig___rarg(x_3, x_4, x_5, x_6, x_146, x_8, x_9); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); +x_147 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_147, 0, x_10); +lean_ctor_set(x_147, 1, x_11); +lean_ctor_set(x_147, 2, x_12); +lean_ctor_set(x_147, 3, x_146); +lean_ctor_set(x_147, 4, x_14); +lean_ctor_set(x_147, 5, x_15); +lean_ctor_set(x_147, 6, x_16); +lean_ctor_set(x_147, 7, x_17); +lean_ctor_set(x_147, 8, x_18); +lean_ctor_set(x_147, 9, x_19); +lean_ctor_set(x_147, 10, x_20); +lean_ctor_set_uint8(x_147, sizeof(void*)*11, x_21); +lean_ctor_set_uint8(x_147, sizeof(void*)*11 + 1, x_22); +x_148 = l_Lean_Meta_Simp_getConfig___rarg(x_3, x_4, x_5, x_6, x_147, x_8, x_9); +x_149 = lean_ctor_get(x_148, 0); lean_inc(x_149); -lean_dec(x_147); -x_150 = lean_st_ref_get(x_4, x_149); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = lean_st_ref_get(x_4, x_150); +x_152 = lean_ctor_get(x_151, 0); lean_inc(x_152); -lean_dec(x_150); -x_153 = lean_ctor_get(x_148, 0); +x_153 = lean_ctor_get(x_151, 1); lean_inc(x_153); -x_154 = lean_ctor_get(x_151, 3); -lean_inc(x_154); lean_dec(x_151); -x_155 = lean_nat_dec_lt(x_153, x_154); +x_154 = lean_ctor_get(x_149, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_152, 3); +lean_inc(x_155); +lean_dec(x_152); +x_156 = lean_nat_dec_lt(x_154, x_155); +lean_dec(x_155); lean_dec(x_154); -lean_dec(x_153); -if (x_155 == 0) +if (x_156 == 0) { -lean_object* x_156; lean_object* x_157; -x_156 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; -x_157 = l_Lean_Core_checkSystem(x_156, x_146, x_8, x_152); -if (lean_obj_tag(x_157) == 0) +lean_object* x_157; lean_object* x_158; +x_157 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; +x_158 = l_Lean_Core_checkSystem(x_157, x_147, x_8, x_153); +if (lean_obj_tag(x_158) == 0) { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_158 = lean_ctor_get(x_157, 1); -lean_inc(x_158); -lean_dec(x_157); -x_159 = lean_st_ref_take(x_4, x_158); -x_160 = lean_ctor_get(x_159, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_159, 1); +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_159 = lean_ctor_get(x_158, 1); +lean_inc(x_159); +lean_dec(x_158); +x_160 = lean_st_ref_take(x_4, x_159); +x_161 = lean_ctor_get(x_160, 0); lean_inc(x_161); -lean_dec(x_159); -x_162 = lean_ctor_get(x_160, 0); +x_162 = lean_ctor_get(x_160, 1); lean_inc(x_162); -x_163 = lean_ctor_get(x_160, 1); +lean_dec(x_160); +x_163 = lean_ctor_get(x_161, 0); lean_inc(x_163); -x_164 = lean_ctor_get(x_160, 2); +x_164 = lean_ctor_get(x_161, 1); lean_inc(x_164); -x_165 = lean_ctor_get(x_160, 3); +x_165 = lean_ctor_get(x_161, 2); lean_inc(x_165); -if (lean_is_exclusive(x_160)) { - lean_ctor_release(x_160, 0); - lean_ctor_release(x_160, 1); - lean_ctor_release(x_160, 2); - lean_ctor_release(x_160, 3); - x_166 = x_160; -} else { - lean_dec_ref(x_160); - x_166 = lean_box(0); +x_166 = lean_ctor_get(x_161, 3); +lean_inc(x_166); +x_167 = lean_ctor_get(x_161, 4); +lean_inc(x_167); +if (lean_is_exclusive(x_161)) { + lean_ctor_release(x_161, 0); + lean_ctor_release(x_161, 1); + lean_ctor_release(x_161, 2); + lean_ctor_release(x_161, 3); + lean_ctor_release(x_161, 4); + x_168 = x_161; +} else { + lean_dec_ref(x_161); + x_168 = lean_box(0); } -x_167 = lean_nat_add(x_165, x_144); -lean_dec(x_165); -if (lean_is_scalar(x_166)) { - x_168 = lean_alloc_ctor(0, 4, 0); +x_169 = lean_nat_add(x_166, x_145); +lean_dec(x_166); +if (lean_is_scalar(x_168)) { + x_170 = lean_alloc_ctor(0, 5, 0); } else { - x_168 = x_166; + x_170 = x_168; } -lean_ctor_set(x_168, 0, x_162); -lean_ctor_set(x_168, 1, x_163); -lean_ctor_set(x_168, 2, x_164); -lean_ctor_set(x_168, 3, x_167); -x_169 = lean_st_ref_set(x_4, x_168, x_161); -x_170 = lean_ctor_get(x_169, 1); -lean_inc(x_170); -lean_dec(x_169); +lean_ctor_set(x_170, 0, x_163); +lean_ctor_set(x_170, 1, x_164); +lean_ctor_set(x_170, 2, x_165); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_167); +x_171 = lean_st_ref_set(x_4, x_170, x_162); +x_172 = lean_ctor_get(x_171, 1); +lean_inc(x_172); +lean_dec(x_171); lean_inc(x_8); -lean_inc(x_146); +lean_inc(x_147); 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_171 = l_Lean_Meta_Simp_pre(x_1, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_170); -if (lean_obj_tag(x_171) == 0) +x_173 = l_Lean_Meta_Simp_pre(x_1, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_172); +if (lean_obj_tag(x_173) == 0) { -lean_object* x_172; -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -switch (lean_obj_tag(x_172)) { +lean_object* x_174; +x_174 = lean_ctor_get(x_173, 0); +lean_inc(x_174); +switch (lean_obj_tag(x_174)) { case 0: { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -lean_dec(x_171); -x_174 = lean_ctor_get(x_172, 0); -lean_inc(x_174); -lean_dec(x_172); -x_175 = l_Lean_Meta_Simp_cacheResult(x_1, x_148, x_174, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_173); +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = lean_ctor_get(x_173, 1); +lean_inc(x_175); +lean_dec(x_173); +x_176 = lean_ctor_get(x_174, 0); +lean_inc(x_176); +lean_dec(x_174); +x_177 = l_Lean_Meta_Simp_cacheResult(x_1, x_149, x_176, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_175); lean_dec(x_8); -lean_dec(x_146); +lean_dec(x_147); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_148); -return x_175; +lean_dec(x_149); +return x_177; } case 1: { -lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_176 = lean_ctor_get(x_171, 1); -lean_inc(x_176); -lean_dec(x_171); -x_177 = lean_ctor_get(x_172, 0); -lean_inc(x_177); -lean_dec(x_172); -x_178 = lean_ctor_get(x_177, 0); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +x_178 = lean_ctor_get(x_173, 1); lean_inc(x_178); +lean_dec(x_173); +x_179 = lean_ctor_get(x_174, 0); +lean_inc(x_179); +lean_dec(x_174); +x_180 = lean_ctor_get(x_179, 0); +lean_inc(x_180); lean_inc(x_8); -lean_inc(x_146); +lean_inc(x_147); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_179 = l_Lean_Meta_Simp_simpLoop(x_178, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_176); -if (lean_obj_tag(x_179) == 0) +x_181 = l_Lean_Meta_Simp_simpLoop(x_180, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_178); +if (lean_obj_tag(x_181) == 0) { -lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_180 = lean_ctor_get(x_179, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); -lean_dec(x_179); +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_ctor_get(x_181, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_181, 1); +lean_inc(x_183); +lean_dec(x_181); lean_inc(x_8); -lean_inc(x_146); +lean_inc(x_147); lean_inc(x_6); lean_inc(x_5); -x_182 = l_Lean_Meta_Simp_Result_mkEqTrans(x_177, x_180, x_5, x_6, x_146, x_8, x_181); -if (lean_obj_tag(x_182) == 0) +x_184 = l_Lean_Meta_Simp_Result_mkEqTrans(x_179, x_182, x_5, x_6, x_147, x_8, x_183); +if (lean_obj_tag(x_184) == 0) { -lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec(x_182); -x_185 = l_Lean_Meta_Simp_cacheResult(x_1, x_148, x_183, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_184); +lean_object* x_185; lean_object* x_186; lean_object* 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_Meta_Simp_cacheResult(x_1, x_149, x_185, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_186); lean_dec(x_8); -lean_dec(x_146); +lean_dec(x_147); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_148); -return x_185; +lean_dec(x_149); +return x_187; } else { -lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_148); -lean_dec(x_146); +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_dec(x_149); +lean_dec(x_147); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -30003,34 +30343,34 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_186 = lean_ctor_get(x_182, 0); -lean_inc(x_186); -x_187 = lean_ctor_get(x_182, 1); -lean_inc(x_187); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_188 = x_182; +x_188 = lean_ctor_get(x_184, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_184, 1); +lean_inc(x_189); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_190 = x_184; } else { - lean_dec_ref(x_182); - x_188 = lean_box(0); + lean_dec_ref(x_184); + x_190 = lean_box(0); } -if (lean_is_scalar(x_188)) { - x_189 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_190)) { + x_191 = lean_alloc_ctor(1, 2, 0); } else { - x_189 = x_188; + x_191 = x_190; } -lean_ctor_set(x_189, 0, x_186); -lean_ctor_set(x_189, 1, x_187); -return x_189; +lean_ctor_set(x_191, 0, x_188); +lean_ctor_set(x_191, 1, x_189); +return x_191; } } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_177); -lean_dec(x_148); -lean_dec(x_146); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_179); +lean_dec(x_149); +lean_dec(x_147); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -30038,72 +30378,72 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_190 = lean_ctor_get(x_179, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_179, 1); -lean_inc(x_191); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_192 = x_179; +x_192 = lean_ctor_get(x_181, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_181, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_181)) { + lean_ctor_release(x_181, 0); + lean_ctor_release(x_181, 1); + x_194 = x_181; } else { - lean_dec_ref(x_179); - x_192 = lean_box(0); + lean_dec_ref(x_181); + x_194 = lean_box(0); } -if (lean_is_scalar(x_192)) { - x_193 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); } else { - x_193 = x_192; + x_195 = x_194; } -lean_ctor_set(x_193, 0, x_190); -lean_ctor_set(x_193, 1, x_191); -return x_193; +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; } } default: { -lean_object* x_194; -x_194 = lean_ctor_get(x_172, 0); -lean_inc(x_194); -lean_dec(x_172); -if (lean_obj_tag(x_194) == 0) +lean_object* x_196; +x_196 = lean_ctor_get(x_174, 0); +lean_inc(x_196); +lean_dec(x_174); +if (lean_obj_tag(x_196) == 0) { -lean_object* x_195; lean_object* x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; -x_195 = lean_ctor_get(x_171, 1); -lean_inc(x_195); -lean_dec(x_171); -x_196 = lean_box(0); -x_197 = 1; +lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_object* x_200; lean_object* x_201; +x_197 = lean_ctor_get(x_173, 1); +lean_inc(x_197); +lean_dec(x_173); +x_198 = lean_box(0); +x_199 = 1; lean_inc(x_1); -x_198 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_198, 0, x_1); -lean_ctor_set(x_198, 1, x_196); -lean_ctor_set_uint8(x_198, sizeof(void*)*2, x_197); -x_199 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_148, x_198, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_195); -lean_dec(x_148); -return x_199; +x_200 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_200, 0, x_1); +lean_ctor_set(x_200, 1, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*2, x_199); +x_201 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_149, x_200, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_197); +lean_dec(x_149); +return x_201; } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; -x_200 = lean_ctor_get(x_171, 1); -lean_inc(x_200); -lean_dec(x_171); -x_201 = lean_ctor_get(x_194, 0); -lean_inc(x_201); -lean_dec(x_194); -x_202 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_148, x_201, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_200); -lean_dec(x_148); -return x_202; +lean_object* x_202; lean_object* x_203; lean_object* x_204; +x_202 = lean_ctor_get(x_173, 1); +lean_inc(x_202); +lean_dec(x_173); +x_203 = lean_ctor_get(x_196, 0); +lean_inc(x_203); +lean_dec(x_196); +x_204 = l_Lean_Meta_Simp_simpLoop_visitPreContinue(x_1, x_149, x_203, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_202); +lean_dec(x_149); +return x_204; } } } } else { -lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -lean_dec(x_148); -lean_dec(x_146); +lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; +lean_dec(x_149); +lean_dec(x_147); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -30111,33 +30451,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_203 = lean_ctor_get(x_171, 0); -lean_inc(x_203); -x_204 = lean_ctor_get(x_171, 1); -lean_inc(x_204); -if (lean_is_exclusive(x_171)) { - lean_ctor_release(x_171, 0); - lean_ctor_release(x_171, 1); - x_205 = x_171; +x_205 = lean_ctor_get(x_173, 0); +lean_inc(x_205); +x_206 = lean_ctor_get(x_173, 1); +lean_inc(x_206); +if (lean_is_exclusive(x_173)) { + lean_ctor_release(x_173, 0); + lean_ctor_release(x_173, 1); + x_207 = x_173; } else { - lean_dec_ref(x_171); - x_205 = lean_box(0); + lean_dec_ref(x_173); + x_207 = lean_box(0); } -if (lean_is_scalar(x_205)) { - x_206 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_207)) { + x_208 = lean_alloc_ctor(1, 2, 0); } else { - x_206 = x_205; + x_208 = x_207; } -lean_ctor_set(x_206, 0, x_203); -lean_ctor_set(x_206, 1, x_204); -return x_206; +lean_ctor_set(x_208, 0, x_205); +lean_ctor_set(x_208, 1, x_206); +return x_208; } } else { -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -lean_dec(x_148); -lean_dec(x_146); +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; +lean_dec(x_149); +lean_dec(x_147); lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -30145,49 +30485,49 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_207 = lean_ctor_get(x_157, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_157, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_209 = x_157; +x_209 = lean_ctor_get(x_158, 0); +lean_inc(x_209); +x_210 = lean_ctor_get(x_158, 1); +lean_inc(x_210); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_211 = x_158; } else { - lean_dec_ref(x_157); - x_209 = lean_box(0); + lean_dec_ref(x_158); + x_211 = lean_box(0); } -if (lean_is_scalar(x_209)) { - x_210 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_211)) { + x_212 = lean_alloc_ctor(1, 2, 0); } else { - x_210 = x_209; + x_212 = x_211; } -lean_ctor_set(x_210, 0, x_207); -lean_ctor_set(x_210, 1, x_208); -return x_210; +lean_ctor_set(x_212, 0, x_209); +lean_ctor_set(x_212, 1, x_210); +return x_212; } } else { -lean_object* x_211; lean_object* x_212; -lean_dec(x_148); +lean_object* x_213; lean_object* x_214; +lean_dec(x_149); lean_dec(x_1); -x_211 = l_Lean_Meta_Simp_simpLoop___closed__2; -x_212 = l_Lean_throwError___at_Lean_Meta_Simp_simpLoop___spec__1(x_211, x_2, x_3, x_4, x_5, x_6, x_146, x_8, x_152); +x_213 = l_Lean_Meta_Simp_simpLoop___closed__2; +x_214 = l_Lean_throwError___at_Lean_Meta_Simp_simpLoop___spec__1(x_213, x_2, x_3, x_4, x_5, x_6, x_147, x_8, x_153); lean_dec(x_8); -lean_dec(x_146); +lean_dec(x_147); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_212; +return x_214; } } } else { -lean_object* x_213; +lean_object* x_215; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -30199,7 +30539,7 @@ lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_1); -x_213 = l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Simp_simpLoop___spec__2(x_15, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_215 = l_Lean_throwMaxRecDepthAt___at_Lean_Meta_Simp_simpLoop___spec__2(x_15, 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); @@ -30207,7 +30547,7 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -return x_213; +return x_215; } } } @@ -31149,7 +31489,7 @@ lean_dec(x_2); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig___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_Meta_Simp_withSimpContext___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: { uint8_t x_8; @@ -31414,39 +31754,24 @@ return x_80; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpContext(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_withSimpConfig___rarg___boxed), 7, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_withSimpContext___rarg___boxed), 7, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_withSimpConfig___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_Meta_Simp_withSimpContext___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_Meta_Simp_withSimpConfig___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Meta_Simp_withSimpContext___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_1, 2); -lean_inc(x_9); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_2); -lean_ctor_set(x_10, 1, x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; -} -} -static lean_object* _init_l_Lean_Meta_Simp_main___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_main_simpMain___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -31456,1258 +31781,1141 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_main___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("numSteps", 8); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Simp_main___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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__1; -x_2 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__2; -x_3 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; -x_4 = l_Lean_Meta_Simp_main___closed__2; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main_simpMain(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { uint8_t x_10; -x_10 = !lean_is_exclusive(x_2); +x_10 = !lean_is_exclusive(x_7); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_11 = lean_ctor_get(x_2, 0); -x_12 = l_Lean_Meta_Simp_Config_updateArith(x_11, x_7, x_8, x_9); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +lean_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; uint8_t x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +x_13 = lean_ctor_get(x_7, 2); +x_14 = lean_ctor_get(x_7, 3); +x_15 = lean_ctor_get(x_7, 4); +x_16 = lean_ctor_get(x_7, 5); +x_17 = lean_ctor_get(x_7, 6); +x_18 = lean_ctor_get(x_7, 7); +x_19 = lean_ctor_get(x_7, 8); +x_20 = lean_ctor_get(x_7, 9); +x_21 = lean_ctor_get(x_7, 10); +x_22 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); +x_23 = 1; +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); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - lean_ctor_release(x_12, 1); - x_15 = x_12; -} else { - lean_dec_ref(x_12); - x_15 = lean_box(0); -} lean_inc(x_13); -lean_ctor_set(x_2, 0, x_13); -x_16 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; -x_17 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; -x_18 = lean_unsigned_to_nat(0u); -x_19 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_19, 0, x_16); -lean_ctor_set(x_19, 1, x_17); -lean_ctor_set(x_19, 2, x_3); -lean_ctor_set(x_19, 3, x_18); -x_20 = !lean_is_exclusive(x_5); -if (x_20 == 0) -{ -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_5, 0); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -uint8_t x_23; uint8_t x_24; uint8_t x_25; -x_23 = lean_ctor_get_uint8(x_13, sizeof(void*)*2 + 6); -lean_dec(x_13); -x_24 = 2; -lean_ctor_set_uint8(x_21, 9, x_24); -lean_ctor_set_uint8(x_21, 11, x_23); -x_25 = !lean_is_exclusive(x_7); -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; uint8_t x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_26 = lean_ctor_get(x_7, 0); -x_27 = lean_ctor_get(x_7, 1); -x_28 = lean_ctor_get(x_7, 2); -x_29 = lean_ctor_get(x_7, 3); -x_30 = lean_ctor_get(x_7, 4); -x_31 = lean_ctor_get(x_7, 5); -x_32 = lean_ctor_get(x_7, 6); -x_33 = lean_ctor_get(x_7, 7); -x_34 = lean_ctor_get(x_7, 8); -x_35 = lean_ctor_get(x_7, 9); -x_36 = lean_ctor_get(x_7, 10); -x_37 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -x_38 = 1; -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_inc(x_27); -lean_inc(x_26); -lean_ctor_set_uint8(x_7, sizeof(void*)*11, x_38); -x_50 = 0; -x_51 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_51, 0, x_26); -lean_ctor_set(x_51, 1, x_27); -lean_ctor_set(x_51, 2, x_28); -lean_ctor_set(x_51, 3, x_29); -lean_ctor_set(x_51, 4, x_30); -lean_ctor_set(x_51, 5, x_31); -lean_ctor_set(x_51, 6, x_32); -lean_ctor_set(x_51, 7, x_33); -lean_ctor_set(x_51, 8, x_34); -lean_ctor_set(x_51, 9, x_35); -lean_ctor_set(x_51, 10, x_36); -lean_ctor_set_uint8(x_51, sizeof(void*)*11, x_50); -lean_ctor_set_uint8(x_51, sizeof(void*)*11 + 1, x_37); -x_52 = lean_st_mk_ref(x_19, x_14); -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_12); +lean_inc(x_11); +lean_ctor_set_uint8(x_7, sizeof(void*)*11, x_23); +x_24 = 0; +x_25 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_25, 0, x_11); +lean_ctor_set(x_25, 1, x_12); +lean_ctor_set(x_25, 2, x_13); +lean_ctor_set(x_25, 3, x_14); +lean_ctor_set(x_25, 4, x_15); +lean_ctor_set(x_25, 5, x_16); +lean_ctor_set(x_25, 6, x_17); +lean_ctor_set(x_25, 7, x_18); +lean_ctor_set(x_25, 8, x_19); +lean_ctor_set(x_25, 9, x_20); +lean_ctor_set(x_25, 10, x_21); +lean_ctor_set_uint8(x_25, sizeof(void*)*11, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*11 + 1, x_22); lean_inc(x_8); -lean_inc(x_51); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_53); -x_55 = lean_simp(x_1, x_4, x_2, x_53, x_5, x_6, x_51, x_8, x_54); -if (lean_obj_tag(x_55) == 0) +lean_inc(x_4); +x_26 = lean_simp(x_1, x_2, x_3, x_4, x_5, x_6, x_25, x_8, x_9); +if (lean_obj_tag(x_26) == 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; uint8_t x_64; +uint8_t x_27; lean_dec(x_7); -lean_dec(x_15); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_55, 1); -lean_inc(x_57); -lean_dec(x_55); -x_58 = lean_st_ref_get(x_53, x_57); -lean_dec(x_53); -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_Meta_Simp_main___closed__3; -x_62 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_61, x_5, x_6, x_51, x_8, x_60); -x_63 = lean_ctor_get(x_62, 0); -lean_inc(x_63); -x_64 = lean_unbox(x_63); -lean_dec(x_63); -if (x_64 == 0) -{ -lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; -x_65 = lean_ctor_get(x_62, 1); -lean_inc(x_65); -lean_dec(x_62); -x_66 = lean_box(0); -x_67 = l_Lean_Meta_Simp_main___lambda__1(x_59, x_56, x_66, x_5, x_6, x_51, x_8, x_65); lean_dec(x_8); -lean_dec(x_51); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_59); -x_68 = !lean_is_exclusive(x_67); -if (x_68 == 0) +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -return x_67; +return x_26; } else { -lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_69 = lean_ctor_get(x_67, 0); -x_70 = lean_ctor_get(x_67, 1); -lean_inc(x_70); -lean_inc(x_69); -lean_dec(x_67); -x_71 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_71, 0, x_69); -lean_ctor_set(x_71, 1, x_70); -return x_71; +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -lean_object* x_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; uint8_t x_84; -x_72 = lean_ctor_get(x_62, 1); -lean_inc(x_72); -lean_dec(x_62); -x_73 = lean_ctor_get(x_59, 3); -lean_inc(x_73); -x_74 = l___private_Init_Data_Repr_0__Nat_reprFast(x_73); -x_75 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_75, 0, x_74); -x_76 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_76, 0, x_75); -x_77 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_76); -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_77); -x_80 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_61, x_79, x_5, x_6, x_51, x_8, x_72); -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 = l_Lean_Meta_Simp_main___lambda__1(x_59, x_56, x_81, x_5, x_6, x_51, x_8, x_82); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_st_ref_get(x_4, x_32); +lean_dec(x_4); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_34, 4); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_37 = l_Lean_Meta_Simp_reportDiag(x_36, x_5, x_6, x_7, x_8, x_35); +if (lean_obj_tag(x_37) == 0) +{ +uint8_t x_38; +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_37, 1); +x_40 = lean_ctor_get(x_37, 0); +lean_dec(x_40); +x_41 = l_Lean_Exception_isRuntime(x_31); +if (x_41 == 0) +{ +lean_dec(x_7); lean_dec(x_8); -lean_dec(x_51); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_81); -lean_dec(x_59); -x_84 = !lean_is_exclusive(x_83); -if (x_84 == 0) +lean_ctor_set_tag(x_37, 1); +lean_ctor_set(x_37, 0, x_31); +return x_37; +} +else { -return x_83; +lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_free_object(x_37); +x_42 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_43 = l_Lean_Meta_throwNestedTacticEx___rarg(x_42, x_31, x_5, x_6, x_7, x_8, x_39); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) +{ +return x_43; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_83, 0); -x_86 = lean_ctor_get(x_83, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_83); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_85); -lean_ctor_set(x_87, 1, x_86); -return x_87; +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_88; lean_object* x_89; -lean_dec(x_53); -lean_dec(x_51); -x_88 = lean_ctor_get(x_55, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_55, 1); -lean_inc(x_89); -lean_dec(x_55); -x_39 = x_88; -x_40 = x_89; -goto block_49; -} -block_49: -{ -uint8_t x_41; -x_41 = l_Lean_Exception_isRuntime(x_39); -if (x_41 == 0) +lean_object* x_48; uint8_t x_49; +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +lean_dec(x_37); +x_49 = l_Lean_Exception_isRuntime(x_31); +if (x_49 == 0) { -lean_object* x_42; +lean_object* x_50; lean_dec(x_7); -lean_dec(x_5); lean_dec(x_8); lean_dec(x_6); -if (lean_is_scalar(x_15)) { - x_42 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_5); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_31); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_51 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_52 = l_Lean_Meta_throwNestedTacticEx___rarg(x_51, x_31, x_5, x_6, x_7, x_8, x_48); +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 { - x_42 = x_15; - lean_ctor_set_tag(x_42, 1); + lean_dec_ref(x_52); + x_55 = lean_box(0); +} +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); +} else { + x_56 = x_55; +} +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; +} } -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; } else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -lean_dec(x_15); -x_43 = l_Lean_Meta_Simp_main___closed__1; -x_44 = l_Lean_Meta_throwNestedTacticEx___rarg(x_43, x_39, x_5, x_6, x_7, x_8, x_40); -x_45 = !lean_is_exclusive(x_44); -if (x_45 == 0) +uint8_t x_57; +lean_dec(x_31); +lean_dec(x_7); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +x_57 = !lean_is_exclusive(x_37); +if (x_57 == 0) { -return x_44; +return x_37; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_44, 0); -x_47 = lean_ctor_get(x_44, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_44); -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_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_37, 0); +x_59 = lean_ctor_get(x_37, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_37); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } } else { -lean_object* x_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; uint8_t x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_90 = lean_ctor_get(x_7, 0); -x_91 = lean_ctor_get(x_7, 1); -x_92 = lean_ctor_get(x_7, 2); -x_93 = lean_ctor_get(x_7, 3); -x_94 = lean_ctor_get(x_7, 4); -x_95 = lean_ctor_get(x_7, 5); -x_96 = lean_ctor_get(x_7, 6); -x_97 = lean_ctor_get(x_7, 7); -x_98 = lean_ctor_get(x_7, 8); -x_99 = lean_ctor_get(x_7, 9); -x_100 = lean_ctor_get(x_7, 10); -x_101 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -lean_inc(x_100); -lean_inc(x_99); -lean_inc(x_98); -lean_inc(x_97); -lean_inc(x_96); -lean_inc(x_95); -lean_inc(x_94); -lean_inc(x_93); -lean_inc(x_92); -lean_inc(x_91); -lean_inc(x_90); +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; uint8_t x_72; uint8_t x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; +x_61 = lean_ctor_get(x_7, 0); +x_62 = lean_ctor_get(x_7, 1); +x_63 = lean_ctor_get(x_7, 2); +x_64 = lean_ctor_get(x_7, 3); +x_65 = lean_ctor_get(x_7, 4); +x_66 = lean_ctor_get(x_7, 5); +x_67 = lean_ctor_get(x_7, 6); +x_68 = lean_ctor_get(x_7, 7); +x_69 = lean_ctor_get(x_7, 8); +x_70 = lean_ctor_get(x_7, 9); +x_71 = lean_ctor_get(x_7, 10); +x_72 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); lean_dec(x_7); -x_102 = 1; -lean_inc(x_100); -lean_inc(x_99); -lean_inc(x_98); -lean_inc(x_97); -lean_inc(x_96); -lean_inc(x_95); -lean_inc(x_94); -lean_inc(x_93); -lean_inc(x_92); -lean_inc(x_91); -lean_inc(x_90); -x_103 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_103, 0, x_90); -lean_ctor_set(x_103, 1, x_91); -lean_ctor_set(x_103, 2, x_92); -lean_ctor_set(x_103, 3, x_93); -lean_ctor_set(x_103, 4, x_94); -lean_ctor_set(x_103, 5, x_95); -lean_ctor_set(x_103, 6, x_96); -lean_ctor_set(x_103, 7, x_97); -lean_ctor_set(x_103, 8, x_98); -lean_ctor_set(x_103, 9, x_99); -lean_ctor_set(x_103, 10, x_100); -lean_ctor_set_uint8(x_103, sizeof(void*)*11, x_102); -lean_ctor_set_uint8(x_103, sizeof(void*)*11 + 1, x_101); -x_115 = 0; -x_116 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_116, 0, x_90); -lean_ctor_set(x_116, 1, x_91); -lean_ctor_set(x_116, 2, x_92); -lean_ctor_set(x_116, 3, x_93); -lean_ctor_set(x_116, 4, x_94); -lean_ctor_set(x_116, 5, x_95); -lean_ctor_set(x_116, 6, x_96); -lean_ctor_set(x_116, 7, x_97); -lean_ctor_set(x_116, 8, x_98); -lean_ctor_set(x_116, 9, x_99); -lean_ctor_set(x_116, 10, x_100); -lean_ctor_set_uint8(x_116, sizeof(void*)*11, x_115); -lean_ctor_set_uint8(x_116, sizeof(void*)*11 + 1, x_101); -x_117 = lean_st_mk_ref(x_19, x_14); -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_73 = 1; +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); +lean_inc(x_68); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_inc(x_62); +lean_inc(x_61); +x_74 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_74, 0, x_61); +lean_ctor_set(x_74, 1, x_62); +lean_ctor_set(x_74, 2, x_63); +lean_ctor_set(x_74, 3, x_64); +lean_ctor_set(x_74, 4, x_65); +lean_ctor_set(x_74, 5, x_66); +lean_ctor_set(x_74, 6, x_67); +lean_ctor_set(x_74, 7, x_68); +lean_ctor_set(x_74, 8, x_69); +lean_ctor_set(x_74, 9, x_70); +lean_ctor_set(x_74, 10, x_71); +lean_ctor_set_uint8(x_74, sizeof(void*)*11, x_73); +lean_ctor_set_uint8(x_74, sizeof(void*)*11 + 1, x_72); +x_75 = 0; +x_76 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_76, 0, x_61); +lean_ctor_set(x_76, 1, x_62); +lean_ctor_set(x_76, 2, x_63); +lean_ctor_set(x_76, 3, x_64); +lean_ctor_set(x_76, 4, x_65); +lean_ctor_set(x_76, 5, x_66); +lean_ctor_set(x_76, 6, x_67); +lean_ctor_set(x_76, 7, x_68); +lean_ctor_set(x_76, 8, x_69); +lean_ctor_set(x_76, 9, x_70); +lean_ctor_set(x_76, 10, x_71); +lean_ctor_set_uint8(x_76, sizeof(void*)*11, x_75); +lean_ctor_set_uint8(x_76, sizeof(void*)*11 + 1, x_72); lean_inc(x_8); -lean_inc(x_116); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_118); -x_120 = lean_simp(x_1, x_4, x_2, x_118, x_5, x_6, x_116, x_8, x_119); -if (lean_obj_tag(x_120) == 0) -{ -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; -lean_dec(x_103); -lean_dec(x_15); -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -lean_dec(x_120); -x_123 = lean_st_ref_get(x_118, x_122); -lean_dec(x_118); -x_124 = lean_ctor_get(x_123, 0); -lean_inc(x_124); -x_125 = lean_ctor_get(x_123, 1); -lean_inc(x_125); -lean_dec(x_123); -x_126 = l_Lean_Meta_Simp_main___closed__3; -x_127 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_126, x_5, x_6, x_116, x_8, x_125); -x_128 = lean_ctor_get(x_127, 0); -lean_inc(x_128); -x_129 = lean_unbox(x_128); -lean_dec(x_128); -if (x_129 == 0) +lean_inc(x_4); +x_77 = lean_simp(x_1, x_2, x_3, x_4, x_5, x_6, x_76, x_8, x_9); +if (lean_obj_tag(x_77) == 0) { -lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_130 = lean_ctor_get(x_127, 1); -lean_inc(x_130); -lean_dec(x_127); -x_131 = lean_box(0); -x_132 = l_Lean_Meta_Simp_main___lambda__1(x_124, x_121, x_131, x_5, x_6, x_116, x_8, x_130); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); lean_dec(x_8); -lean_dec(x_116); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_124); -x_133 = lean_ctor_get(x_132, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_132, 1); -lean_inc(x_134); -if (lean_is_exclusive(x_132)) { - lean_ctor_release(x_132, 0); - lean_ctor_release(x_132, 1); - x_135 = x_132; +lean_dec(x_4); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_80 = x_77; } else { - lean_dec_ref(x_132); - x_135 = lean_box(0); + lean_dec_ref(x_77); + x_80 = lean_box(0); } -if (lean_is_scalar(x_135)) { - x_136 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_136 = x_135; + x_81 = x_80; } -lean_ctor_set(x_136, 0, x_133); -lean_ctor_set(x_136, 1, x_134); -return x_136; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_137 = lean_ctor_get(x_127, 1); -lean_inc(x_137); -lean_dec(x_127); -x_138 = lean_ctor_get(x_124, 3); -lean_inc(x_138); -x_139 = l___private_Init_Data_Repr_0__Nat_reprFast(x_138); -x_140 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_140, 0, x_139); -x_141 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_141, 0, x_140); -x_142 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_143 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_141); -x_144 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_142); -x_145 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_126, x_144, x_5, x_6, x_116, x_8, x_137); -x_146 = lean_ctor_get(x_145, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_145, 1); -lean_inc(x_147); -lean_dec(x_145); -x_148 = l_Lean_Meta_Simp_main___lambda__1(x_124, x_121, x_146, x_5, x_6, x_116, x_8, x_147); +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_82 = lean_ctor_get(x_77, 0); +lean_inc(x_82); +x_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +lean_dec(x_77); +x_84 = lean_st_ref_get(x_4, x_83); +lean_dec(x_4); +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, 4); +lean_inc(x_87); +lean_dec(x_85); +lean_inc(x_8); +lean_inc(x_74); +lean_inc(x_6); +lean_inc(x_5); +x_88 = l_Lean_Meta_Simp_reportDiag(x_87, x_5, x_6, x_74, x_8, x_86); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_ctor_get(x_88, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_90 = x_88; +} else { + lean_dec_ref(x_88); + x_90 = lean_box(0); +} +x_91 = l_Lean_Exception_isRuntime(x_82); +if (x_91 == 0) +{ +lean_object* x_92; +lean_dec(x_74); lean_dec(x_8); -lean_dec(x_116); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_146); -lean_dec(x_124); -x_149 = lean_ctor_get(x_148, 0); -lean_inc(x_149); -x_150 = lean_ctor_get(x_148, 1); -lean_inc(x_150); -if (lean_is_exclusive(x_148)) { - lean_ctor_release(x_148, 0); - lean_ctor_release(x_148, 1); - x_151 = x_148; +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_148); - x_151 = lean_box(0); + x_92 = x_90; + lean_ctor_set_tag(x_92, 1); } -if (lean_is_scalar(x_151)) { - x_152 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_82); +lean_ctor_set(x_92, 1, x_89); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_90); +x_93 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_94 = l_Lean_Meta_throwNestedTacticEx___rarg(x_93, x_82, x_5, x_6, x_74, x_8, x_89); +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); +lean_inc(x_96); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_97 = x_94; } else { - x_152 = x_151; + lean_dec_ref(x_94); + x_97 = lean_box(0); } -lean_ctor_set(x_152, 0, x_149); -lean_ctor_set(x_152, 1, x_150); -return x_152; +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_153; lean_object* x_154; -lean_dec(x_118); -lean_dec(x_116); -x_153 = lean_ctor_get(x_120, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_120, 1); -lean_inc(x_154); -lean_dec(x_120); -x_104 = x_153; -x_105 = x_154; -goto block_114; } -block_114: -{ -uint8_t x_106; -x_106 = l_Lean_Exception_isRuntime(x_104); -if (x_106 == 0) +else { -lean_object* x_107; -lean_dec(x_103); -lean_dec(x_5); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_82); +lean_dec(x_74); lean_dec(x_8); lean_dec(x_6); -if (lean_is_scalar(x_15)) { - x_107 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_5); +x_99 = lean_ctor_get(x_88, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_88, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_101 = x_88; } else { - x_107 = x_15; - lean_ctor_set_tag(x_107, 1); + lean_dec_ref(x_88); + x_101 = lean_box(0); } -lean_ctor_set(x_107, 0, x_104); -lean_ctor_set(x_107, 1, x_105); -return x_107; +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); +} else { + x_102 = x_101; } -else +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; -lean_dec(x_15); -x_108 = l_Lean_Meta_Simp_main___closed__1; -x_109 = l_Lean_Meta_throwNestedTacticEx___rarg(x_108, x_104, x_5, x_6, x_103, x_8, x_105); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); -lean_inc(x_111); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - x_112 = x_109; -} else { - lean_dec_ref(x_109); - x_112 = lean_box(0); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = lean_ctor_get(x_1, 2); +x_10 = lean_ctor_get(x_1, 4); +lean_inc(x_10); +lean_inc(x_9); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_2); +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_8); +return x_13; } -if (lean_is_scalar(x_112)) { - x_113 = lean_alloc_ctor(1, 2, 0); -} else { - x_113 = x_112; } -lean_ctor_set(x_113, 0, x_110); -lean_ctor_set(x_113, 1, x_111); -return x_113; +static lean_object* _init_l_Lean_Meta_Simp_main___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("numSteps", 8); +return x_1; } } +static lean_object* _init_l_Lean_Meta_Simp_main___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___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__1; +x_2 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__2; +x_3 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; +x_4 = l_Lean_Meta_Simp_main___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_155; uint8_t x_156; uint8_t 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; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; uint8_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -x_155 = lean_ctor_get_uint8(x_21, 0); -x_156 = lean_ctor_get_uint8(x_21, 1); -x_157 = lean_ctor_get_uint8(x_21, 2); -x_158 = lean_ctor_get_uint8(x_21, 3); -x_159 = lean_ctor_get_uint8(x_21, 4); -x_160 = lean_ctor_get_uint8(x_21, 5); -x_161 = lean_ctor_get_uint8(x_21, 6); -x_162 = lean_ctor_get_uint8(x_21, 7); -x_163 = lean_ctor_get_uint8(x_21, 8); -x_164 = lean_ctor_get_uint8(x_21, 10); -x_165 = lean_ctor_get_uint8(x_21, 12); -lean_dec(x_21); -x_166 = lean_ctor_get_uint8(x_13, sizeof(void*)*2 + 6); +uint8_t x_10; +x_10 = !lean_is_exclusive(x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_11 = lean_ctor_get(x_2, 0); +x_12 = lean_ctor_get(x_2, 4); +lean_dec(x_12); +x_13 = l_Lean_Meta_Simp_Config_updateArith(x_11, x_7, x_8, x_9); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); lean_dec(x_13); -x_167 = 2; -x_168 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_168, 0, x_155); -lean_ctor_set_uint8(x_168, 1, x_156); -lean_ctor_set_uint8(x_168, 2, x_157); -lean_ctor_set_uint8(x_168, 3, x_158); -lean_ctor_set_uint8(x_168, 4, x_159); -lean_ctor_set_uint8(x_168, 5, x_160); -lean_ctor_set_uint8(x_168, 6, x_161); -lean_ctor_set_uint8(x_168, 7, x_162); -lean_ctor_set_uint8(x_168, 8, x_163); -lean_ctor_set_uint8(x_168, 9, x_167); -lean_ctor_set_uint8(x_168, 10, x_164); -lean_ctor_set_uint8(x_168, 11, x_166); -lean_ctor_set_uint8(x_168, 12, x_165); -lean_ctor_set(x_5, 0, x_168); -x_169 = lean_ctor_get(x_7, 0); -lean_inc(x_169); -x_170 = lean_ctor_get(x_7, 1); -lean_inc(x_170); -x_171 = lean_ctor_get(x_7, 2); -lean_inc(x_171); -x_172 = lean_ctor_get(x_7, 3); -lean_inc(x_172); -x_173 = lean_ctor_get(x_7, 4); -lean_inc(x_173); -x_174 = lean_ctor_get(x_7, 5); -lean_inc(x_174); -x_175 = lean_ctor_get(x_7, 6); -lean_inc(x_175); -x_176 = lean_ctor_get(x_7, 7); -lean_inc(x_176); -x_177 = lean_ctor_get(x_7, 8); -lean_inc(x_177); -x_178 = lean_ctor_get(x_7, 9); -lean_inc(x_178); -x_179 = lean_ctor_get(x_7, 10); -lean_inc(x_179); -x_180 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - lean_ctor_release(x_7, 7); - lean_ctor_release(x_7, 8); - lean_ctor_release(x_7, 9); - lean_ctor_release(x_7, 10); - x_181 = x_7; -} else { - lean_dec_ref(x_7); - x_181 = lean_box(0); -} -x_182 = 1; -lean_inc(x_179); -lean_inc(x_178); -lean_inc(x_177); -lean_inc(x_176); -lean_inc(x_175); -lean_inc(x_174); -lean_inc(x_173); -lean_inc(x_172); -lean_inc(x_171); -lean_inc(x_170); -lean_inc(x_169); -if (lean_is_scalar(x_181)) { - x_183 = lean_alloc_ctor(0, 11, 2); -} else { - x_183 = x_181; -} -lean_ctor_set(x_183, 0, x_169); -lean_ctor_set(x_183, 1, x_170); -lean_ctor_set(x_183, 2, x_171); -lean_ctor_set(x_183, 3, x_172); -lean_ctor_set(x_183, 4, x_173); -lean_ctor_set(x_183, 5, x_174); -lean_ctor_set(x_183, 6, x_175); -lean_ctor_set(x_183, 7, x_176); -lean_ctor_set(x_183, 8, x_177); -lean_ctor_set(x_183, 9, x_178); -lean_ctor_set(x_183, 10, x_179); -lean_ctor_set_uint8(x_183, sizeof(void*)*11, x_182); -lean_ctor_set_uint8(x_183, sizeof(void*)*11 + 1, x_180); -x_195 = 0; -x_196 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_196, 0, x_169); -lean_ctor_set(x_196, 1, x_170); -lean_ctor_set(x_196, 2, x_171); -lean_ctor_set(x_196, 3, x_172); -lean_ctor_set(x_196, 4, x_173); -lean_ctor_set(x_196, 5, x_174); -lean_ctor_set(x_196, 6, x_175); -lean_ctor_set(x_196, 7, x_176); -lean_ctor_set(x_196, 8, x_177); -lean_ctor_set(x_196, 9, x_178); -lean_ctor_set(x_196, 10, x_179); -lean_ctor_set_uint8(x_196, sizeof(void*)*11, x_195); -lean_ctor_set_uint8(x_196, sizeof(void*)*11 + 1, x_180); -x_197 = lean_st_mk_ref(x_19, x_14); -x_198 = lean_ctor_get(x_197, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_197, 1); -lean_inc(x_199); -lean_dec(x_197); +x_16 = !lean_is_exclusive(x_5); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_17 = lean_ctor_get(x_5, 0); +x_18 = lean_ctor_get(x_5, 1); +lean_inc(x_18); +x_19 = lean_local_ctx_num_indices(x_18); +lean_inc(x_14); +lean_ctor_set(x_2, 4, x_19); +lean_ctor_set(x_2, 0, x_14); +x_20 = lean_ctor_get(x_3, 0); +x_21 = lean_ctor_get(x_3, 1); +x_22 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_23 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; +x_24 = lean_unsigned_to_nat(0u); +lean_inc(x_21); +lean_inc(x_20); +x_25 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_25, 0, x_22); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_24); +lean_ctor_set(x_25, 4, x_21); +x_26 = !lean_is_exclusive(x_17); +if (x_26 == 0) +{ +uint8_t x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_27 = lean_ctor_get_uint8(x_14, sizeof(void*)*2 + 6); +lean_dec(x_14); +x_28 = 2; +lean_ctor_set_uint8(x_17, 9, x_28); +lean_ctor_set_uint8(x_17, 11, x_27); +x_29 = lean_st_mk_ref(x_25, x_15); +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_8); -lean_inc(x_196); +lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_198); -x_200 = lean_simp(x_1, x_4, x_2, x_198, x_5, x_6, x_196, x_8, x_199); -if (lean_obj_tag(x_200) == 0) +lean_inc(x_30); +x_32 = l_Lean_Meta_Simp_main_simpMain(x_1, x_4, x_2, x_30, x_5, x_6, x_7, x_8, x_31); +if (lean_obj_tag(x_32) == 0) { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -lean_dec(x_183); -lean_dec(x_15); -x_201 = lean_ctor_get(x_200, 0); -lean_inc(x_201); -x_202 = lean_ctor_get(x_200, 1); -lean_inc(x_202); -lean_dec(x_200); -x_203 = lean_st_ref_get(x_198, x_202); -lean_dec(x_198); -x_204 = lean_ctor_get(x_203, 0); -lean_inc(x_204); -x_205 = lean_ctor_get(x_203, 1); -lean_inc(x_205); -lean_dec(x_203); -x_206 = l_Lean_Meta_Simp_main___closed__3; -x_207 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_206, x_5, x_6, x_196, x_8, x_205); -x_208 = lean_ctor_get(x_207, 0); -lean_inc(x_208); -x_209 = lean_unbox(x_208); -lean_dec(x_208); -if (x_209 == 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; uint8_t x_41; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = lean_st_ref_get(x_30, x_34); +lean_dec(x_30); +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_Meta_Simp_main___closed__2; +x_39 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_38, x_5, x_6, x_7, x_8, x_37); +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 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; -x_210 = lean_ctor_get(x_207, 1); -lean_inc(x_210); -lean_dec(x_207); -x_211 = lean_box(0); -x_212 = l_Lean_Meta_Simp_main___lambda__1(x_204, x_201, x_211, x_5, x_6, x_196, x_8, x_210); +lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; +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_Meta_Simp_main___lambda__1(x_36, x_33, x_43, x_5, x_6, x_7, x_8, x_42); lean_dec(x_8); -lean_dec(x_196); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_204); -x_213 = lean_ctor_get(x_212, 0); -lean_inc(x_213); -x_214 = lean_ctor_get(x_212, 1); -lean_inc(x_214); -if (lean_is_exclusive(x_212)) { - lean_ctor_release(x_212, 0); - lean_ctor_release(x_212, 1); - x_215 = x_212; -} else { - lean_dec_ref(x_212); - x_215 = lean_box(0); +lean_dec(x_36); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +return x_44; } -if (lean_is_scalar(x_215)) { - x_216 = lean_alloc_ctor(0, 2, 0); -} else { - x_216 = x_215; +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_44); +x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_48, 0, x_46); +lean_ctor_set(x_48, 1, x_47); +return x_48; } -lean_ctor_set(x_216, 0, x_213); -lean_ctor_set(x_216, 1, x_214); -return x_216; } else { -lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; -x_217 = lean_ctor_get(x_207, 1); -lean_inc(x_217); -lean_dec(x_207); -x_218 = lean_ctor_get(x_204, 3); -lean_inc(x_218); -x_219 = l___private_Init_Data_Repr_0__Nat_reprFast(x_218); -x_220 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_220, 0, x_219); -x_221 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_221, 0, x_220); -x_222 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_223 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_223, 0, x_222); -lean_ctor_set(x_223, 1, x_221); -x_224 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_222); -x_225 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_206, x_224, x_5, x_6, x_196, x_8, x_217); -x_226 = lean_ctor_get(x_225, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_225, 1); -lean_inc(x_227); -lean_dec(x_225); -x_228 = l_Lean_Meta_Simp_main___lambda__1(x_204, x_201, x_226, x_5, x_6, x_196, x_8, x_227); +lean_object* x_49; 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; uint8_t x_61; +x_49 = lean_ctor_get(x_39, 1); +lean_inc(x_49); +lean_dec(x_39); +x_50 = lean_ctor_get(x_36, 3); +lean_inc(x_50); +x_51 = l___private_Init_Data_Repr_0__Nat_reprFast(x_50); +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); +x_54 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +x_55 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_55, 0, x_54); +lean_ctor_set(x_55, 1, x_53); +x_56 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_54); +x_57 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_38, x_56, x_5, x_6, x_7, x_8, x_49); +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_Simp_main___lambda__1(x_36, x_33, x_58, x_5, x_6, x_7, x_8, x_59); lean_dec(x_8); -lean_dec(x_196); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_226); -lean_dec(x_204); -x_229 = lean_ctor_get(x_228, 0); -lean_inc(x_229); -x_230 = lean_ctor_get(x_228, 1); -lean_inc(x_230); -if (lean_is_exclusive(x_228)) { - lean_ctor_release(x_228, 0); - lean_ctor_release(x_228, 1); - x_231 = x_228; -} else { - lean_dec_ref(x_228); - x_231 = lean_box(0); +lean_dec(x_58); +lean_dec(x_36); +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +return x_60; } -if (lean_is_scalar(x_231)) { - x_232 = lean_alloc_ctor(0, 2, 0); -} else { - x_232 = x_231; +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(0, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } -lean_ctor_set(x_232, 0, x_229); -lean_ctor_set(x_232, 1, x_230); -return x_232; } } else { -lean_object* x_233; lean_object* x_234; -lean_dec(x_198); -lean_dec(x_196); -x_233 = lean_ctor_get(x_200, 0); -lean_inc(x_233); -x_234 = lean_ctor_get(x_200, 1); -lean_inc(x_234); -lean_dec(x_200); -x_184 = x_233; -x_185 = x_234; -goto block_194; +uint8_t x_65; +lean_dec(x_30); +lean_dec(x_5); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_65 = !lean_is_exclusive(x_32); +if (x_65 == 0) +{ +return x_32; } -block_194: +else { -uint8_t x_186; -x_186 = l_Lean_Exception_isRuntime(x_184); -if (x_186 == 0) +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_32, 0); +x_67 = lean_ctor_get(x_32, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_32); +x_68 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_68, 0, x_66); +lean_ctor_set(x_68, 1, x_67); +return x_68; +} +} +} +else { -lean_object* x_187; -lean_dec(x_183); -lean_dec(x_5); +uint8_t x_69; uint8_t x_70; uint8_t x_71; uint8_t x_72; uint8_t x_73; uint8_t x_74; uint8_t x_75; uint8_t x_76; uint8_t x_77; uint8_t x_78; uint8_t x_79; uint8_t 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; +x_69 = lean_ctor_get_uint8(x_17, 0); +x_70 = lean_ctor_get_uint8(x_17, 1); +x_71 = lean_ctor_get_uint8(x_17, 2); +x_72 = lean_ctor_get_uint8(x_17, 3); +x_73 = lean_ctor_get_uint8(x_17, 4); +x_74 = lean_ctor_get_uint8(x_17, 5); +x_75 = lean_ctor_get_uint8(x_17, 6); +x_76 = lean_ctor_get_uint8(x_17, 7); +x_77 = lean_ctor_get_uint8(x_17, 8); +x_78 = lean_ctor_get_uint8(x_17, 10); +x_79 = lean_ctor_get_uint8(x_17, 12); +lean_dec(x_17); +x_80 = lean_ctor_get_uint8(x_14, sizeof(void*)*2 + 6); +lean_dec(x_14); +x_81 = 2; +x_82 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_82, 0, x_69); +lean_ctor_set_uint8(x_82, 1, x_70); +lean_ctor_set_uint8(x_82, 2, x_71); +lean_ctor_set_uint8(x_82, 3, x_72); +lean_ctor_set_uint8(x_82, 4, x_73); +lean_ctor_set_uint8(x_82, 5, x_74); +lean_ctor_set_uint8(x_82, 6, x_75); +lean_ctor_set_uint8(x_82, 7, x_76); +lean_ctor_set_uint8(x_82, 8, x_77); +lean_ctor_set_uint8(x_82, 9, x_81); +lean_ctor_set_uint8(x_82, 10, x_78); +lean_ctor_set_uint8(x_82, 11, x_80); +lean_ctor_set_uint8(x_82, 12, x_79); +lean_ctor_set(x_5, 0, x_82); +x_83 = lean_st_mk_ref(x_25, x_15); +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); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_84); +x_86 = l_Lean_Meta_Simp_main_simpMain(x_1, x_4, x_2, x_84, x_5, x_6, x_7, x_8, x_85); +if (lean_obj_tag(x_86) == 0) +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +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_st_ref_get(x_84, x_88); +lean_dec(x_84); +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 = l_Lean_Meta_Simp_main___closed__2; +x_93 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_92, x_5, x_6, x_7, x_8, x_91); +x_94 = lean_ctor_get(x_93, 0); +lean_inc(x_94); +x_95 = lean_unbox(x_94); +lean_dec(x_94); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_96 = lean_ctor_get(x_93, 1); +lean_inc(x_96); +lean_dec(x_93); +x_97 = lean_box(0); +x_98 = l_Lean_Meta_Simp_main___lambda__1(x_90, x_87, x_97, x_5, x_6, x_7, x_8, x_96); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -if (lean_is_scalar(x_15)) { - x_187 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_5); +lean_dec(x_90); +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_98, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_98)) { + lean_ctor_release(x_98, 0); + lean_ctor_release(x_98, 1); + x_101 = x_98; } else { - x_187 = x_15; - lean_ctor_set_tag(x_187, 1); + lean_dec_ref(x_98); + x_101 = lean_box(0); } -lean_ctor_set(x_187, 0, x_184); -lean_ctor_set(x_187, 1, x_185); -return x_187; +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(0, 2, 0); +} else { + x_102 = x_101; +} +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; -lean_dec(x_15); -x_188 = l_Lean_Meta_Simp_main___closed__1; -x_189 = l_Lean_Meta_throwNestedTacticEx___rarg(x_188, x_184, x_5, x_6, x_183, x_8, x_185); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -if (lean_is_exclusive(x_189)) { - lean_ctor_release(x_189, 0); - lean_ctor_release(x_189, 1); - x_192 = x_189; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_103 = lean_ctor_get(x_93, 1); +lean_inc(x_103); +lean_dec(x_93); +x_104 = lean_ctor_get(x_90, 3); +lean_inc(x_104); +x_105 = l___private_Init_Data_Repr_0__Nat_reprFast(x_104); +x_106 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_107 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_107, 0, x_106); +x_108 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +x_109 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_109, 0, x_108); +lean_ctor_set(x_109, 1, x_107); +x_110 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_110, 0, x_109); +lean_ctor_set(x_110, 1, x_108); +x_111 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_92, x_110, x_5, x_6, x_7, x_8, x_103); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +lean_dec(x_111); +x_114 = l_Lean_Meta_Simp_main___lambda__1(x_90, x_87, x_112, x_5, x_6, x_7, x_8, x_113); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_112); +lean_dec(x_90); +x_115 = lean_ctor_get(x_114, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_114, 1); +lean_inc(x_116); +if (lean_is_exclusive(x_114)) { + lean_ctor_release(x_114, 0); + lean_ctor_release(x_114, 1); + x_117 = x_114; } else { - lean_dec_ref(x_189); - x_192 = lean_box(0); + lean_dec_ref(x_114); + x_117 = lean_box(0); } -if (lean_is_scalar(x_192)) { - x_193 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_117)) { + x_118 = lean_alloc_ctor(0, 2, 0); } else { - x_193 = x_192; -} -lean_ctor_set(x_193, 0, x_190); -lean_ctor_set(x_193, 1, x_191); -return x_193; -} + x_118 = x_117; } +lean_ctor_set(x_118, 0, x_115); +lean_ctor_set(x_118, 1, x_116); +return x_118; } } else { -lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; uint8_t x_241; uint8_t x_242; uint8_t x_243; uint8_t x_244; uint8_t x_245; uint8_t x_246; uint8_t x_247; uint8_t x_248; uint8_t x_249; uint8_t x_250; uint8_t x_251; uint8_t x_252; uint8_t x_253; lean_object* x_254; uint8_t x_255; uint8_t x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; uint8_t x_270; lean_object* x_271; uint8_t x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; uint8_t x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; -x_235 = lean_ctor_get(x_5, 0); -x_236 = lean_ctor_get(x_5, 1); -x_237 = lean_ctor_get(x_5, 2); -x_238 = lean_ctor_get(x_5, 3); -x_239 = lean_ctor_get(x_5, 4); -x_240 = lean_ctor_get(x_5, 5); -x_241 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); -x_242 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); -lean_inc(x_240); -lean_inc(x_239); -lean_inc(x_238); -lean_inc(x_237); -lean_inc(x_236); -lean_inc(x_235); +lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +lean_dec(x_84); lean_dec(x_5); -x_243 = lean_ctor_get_uint8(x_235, 0); -x_244 = lean_ctor_get_uint8(x_235, 1); -x_245 = lean_ctor_get_uint8(x_235, 2); -x_246 = lean_ctor_get_uint8(x_235, 3); -x_247 = lean_ctor_get_uint8(x_235, 4); -x_248 = lean_ctor_get_uint8(x_235, 5); -x_249 = lean_ctor_get_uint8(x_235, 6); -x_250 = lean_ctor_get_uint8(x_235, 7); -x_251 = lean_ctor_get_uint8(x_235, 8); -x_252 = lean_ctor_get_uint8(x_235, 10); -x_253 = lean_ctor_get_uint8(x_235, 12); -if (lean_is_exclusive(x_235)) { - x_254 = x_235; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_119 = lean_ctor_get(x_86, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_86, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_86)) { + lean_ctor_release(x_86, 0); + lean_ctor_release(x_86, 1); + x_121 = x_86; } else { - lean_dec_ref(x_235); - x_254 = lean_box(0); + lean_dec_ref(x_86); + x_121 = lean_box(0); } -x_255 = lean_ctor_get_uint8(x_13, sizeof(void*)*2 + 6); -lean_dec(x_13); -x_256 = 2; -if (lean_is_scalar(x_254)) { - x_257 = lean_alloc_ctor(0, 0, 13); -} else { - x_257 = x_254; -} -lean_ctor_set_uint8(x_257, 0, x_243); -lean_ctor_set_uint8(x_257, 1, x_244); -lean_ctor_set_uint8(x_257, 2, x_245); -lean_ctor_set_uint8(x_257, 3, x_246); -lean_ctor_set_uint8(x_257, 4, x_247); -lean_ctor_set_uint8(x_257, 5, x_248); -lean_ctor_set_uint8(x_257, 6, x_249); -lean_ctor_set_uint8(x_257, 7, x_250); -lean_ctor_set_uint8(x_257, 8, x_251); -lean_ctor_set_uint8(x_257, 9, x_256); -lean_ctor_set_uint8(x_257, 10, x_252); -lean_ctor_set_uint8(x_257, 11, x_255); -lean_ctor_set_uint8(x_257, 12, x_253); -x_258 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_258, 0, x_257); -lean_ctor_set(x_258, 1, x_236); -lean_ctor_set(x_258, 2, x_237); -lean_ctor_set(x_258, 3, x_238); -lean_ctor_set(x_258, 4, x_239); -lean_ctor_set(x_258, 5, x_240); -lean_ctor_set_uint8(x_258, sizeof(void*)*6, x_241); -lean_ctor_set_uint8(x_258, sizeof(void*)*6 + 1, x_242); -x_259 = lean_ctor_get(x_7, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_7, 1); -lean_inc(x_260); -x_261 = lean_ctor_get(x_7, 2); -lean_inc(x_261); -x_262 = lean_ctor_get(x_7, 3); -lean_inc(x_262); -x_263 = lean_ctor_get(x_7, 4); -lean_inc(x_263); -x_264 = lean_ctor_get(x_7, 5); -lean_inc(x_264); -x_265 = lean_ctor_get(x_7, 6); -lean_inc(x_265); -x_266 = lean_ctor_get(x_7, 7); -lean_inc(x_266); -x_267 = lean_ctor_get(x_7, 8); -lean_inc(x_267); -x_268 = lean_ctor_get(x_7, 9); -lean_inc(x_268); -x_269 = lean_ctor_get(x_7, 10); -lean_inc(x_269); -x_270 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - lean_ctor_release(x_7, 7); - lean_ctor_release(x_7, 8); - lean_ctor_release(x_7, 9); - lean_ctor_release(x_7, 10); - x_271 = x_7; +if (lean_is_scalar(x_121)) { + x_122 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_7); - x_271 = lean_box(0); + x_122 = x_121; } -x_272 = 1; -lean_inc(x_269); -lean_inc(x_268); -lean_inc(x_267); -lean_inc(x_266); -lean_inc(x_265); -lean_inc(x_264); -lean_inc(x_263); -lean_inc(x_262); -lean_inc(x_261); -lean_inc(x_260); -lean_inc(x_259); -if (lean_is_scalar(x_271)) { - x_273 = lean_alloc_ctor(0, 11, 2); -} else { - x_273 = x_271; +lean_ctor_set(x_122, 0, x_119); +lean_ctor_set(x_122, 1, x_120); +return x_122; } -lean_ctor_set(x_273, 0, x_259); -lean_ctor_set(x_273, 1, x_260); -lean_ctor_set(x_273, 2, x_261); -lean_ctor_set(x_273, 3, x_262); -lean_ctor_set(x_273, 4, x_263); -lean_ctor_set(x_273, 5, x_264); -lean_ctor_set(x_273, 6, x_265); -lean_ctor_set(x_273, 7, x_266); -lean_ctor_set(x_273, 8, x_267); -lean_ctor_set(x_273, 9, x_268); -lean_ctor_set(x_273, 10, x_269); -lean_ctor_set_uint8(x_273, sizeof(void*)*11, x_272); -lean_ctor_set_uint8(x_273, sizeof(void*)*11 + 1, x_270); -x_285 = 0; -x_286 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_286, 0, x_259); -lean_ctor_set(x_286, 1, x_260); -lean_ctor_set(x_286, 2, x_261); -lean_ctor_set(x_286, 3, x_262); -lean_ctor_set(x_286, 4, x_263); -lean_ctor_set(x_286, 5, x_264); -lean_ctor_set(x_286, 6, x_265); -lean_ctor_set(x_286, 7, x_266); -lean_ctor_set(x_286, 8, x_267); -lean_ctor_set(x_286, 9, x_268); -lean_ctor_set(x_286, 10, x_269); -lean_ctor_set_uint8(x_286, sizeof(void*)*11, x_285); -lean_ctor_set_uint8(x_286, sizeof(void*)*11 + 1, x_270); -x_287 = lean_st_mk_ref(x_19, x_14); -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); +} +} +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; uint8_t x_129; uint8_t 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; uint8_t x_138; uint8_t x_139; uint8_t x_140; uint8_t x_141; uint8_t x_142; uint8_t x_143; uint8_t x_144; uint8_t x_145; uint8_t x_146; uint8_t x_147; uint8_t x_148; lean_object* x_149; uint8_t x_150; uint8_t x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +x_123 = lean_ctor_get(x_5, 0); +x_124 = lean_ctor_get(x_5, 1); +x_125 = lean_ctor_get(x_5, 2); +x_126 = lean_ctor_get(x_5, 3); +x_127 = lean_ctor_get(x_5, 4); +x_128 = lean_ctor_get(x_5, 5); +x_129 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); +x_130 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); +lean_inc(x_128); +lean_inc(x_127); +lean_inc(x_126); +lean_inc(x_125); +lean_inc(x_124); +lean_inc(x_123); +lean_dec(x_5); +lean_inc(x_124); +x_131 = lean_local_ctx_num_indices(x_124); +lean_inc(x_14); +lean_ctor_set(x_2, 4, x_131); +lean_ctor_set(x_2, 0, x_14); +x_132 = lean_ctor_get(x_3, 0); +x_133 = lean_ctor_get(x_3, 1); +x_134 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_135 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; +x_136 = lean_unsigned_to_nat(0u); +lean_inc(x_133); +lean_inc(x_132); +x_137 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_137, 0, x_134); +lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 2, x_132); +lean_ctor_set(x_137, 3, x_136); +lean_ctor_set(x_137, 4, x_133); +x_138 = lean_ctor_get_uint8(x_123, 0); +x_139 = lean_ctor_get_uint8(x_123, 1); +x_140 = lean_ctor_get_uint8(x_123, 2); +x_141 = lean_ctor_get_uint8(x_123, 3); +x_142 = lean_ctor_get_uint8(x_123, 4); +x_143 = lean_ctor_get_uint8(x_123, 5); +x_144 = lean_ctor_get_uint8(x_123, 6); +x_145 = lean_ctor_get_uint8(x_123, 7); +x_146 = lean_ctor_get_uint8(x_123, 8); +x_147 = lean_ctor_get_uint8(x_123, 10); +x_148 = lean_ctor_get_uint8(x_123, 12); +if (lean_is_exclusive(x_123)) { + x_149 = x_123; +} else { + lean_dec_ref(x_123); + x_149 = lean_box(0); +} +x_150 = lean_ctor_get_uint8(x_14, sizeof(void*)*2 + 6); +lean_dec(x_14); +x_151 = 2; +if (lean_is_scalar(x_149)) { + x_152 = lean_alloc_ctor(0, 0, 13); +} else { + x_152 = x_149; +} +lean_ctor_set_uint8(x_152, 0, x_138); +lean_ctor_set_uint8(x_152, 1, x_139); +lean_ctor_set_uint8(x_152, 2, x_140); +lean_ctor_set_uint8(x_152, 3, x_141); +lean_ctor_set_uint8(x_152, 4, x_142); +lean_ctor_set_uint8(x_152, 5, x_143); +lean_ctor_set_uint8(x_152, 6, x_144); +lean_ctor_set_uint8(x_152, 7, x_145); +lean_ctor_set_uint8(x_152, 8, x_146); +lean_ctor_set_uint8(x_152, 9, x_151); +lean_ctor_set_uint8(x_152, 10, x_147); +lean_ctor_set_uint8(x_152, 11, x_150); +lean_ctor_set_uint8(x_152, 12, x_148); +x_153 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_153, 0, x_152); +lean_ctor_set(x_153, 1, x_124); +lean_ctor_set(x_153, 2, x_125); +lean_ctor_set(x_153, 3, x_126); +lean_ctor_set(x_153, 4, x_127); +lean_ctor_set(x_153, 5, x_128); +lean_ctor_set_uint8(x_153, sizeof(void*)*6, x_129); +lean_ctor_set_uint8(x_153, sizeof(void*)*6 + 1, x_130); +x_154 = lean_st_mk_ref(x_137, x_15); +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +lean_dec(x_154); lean_inc(x_8); -lean_inc(x_286); +lean_inc(x_7); lean_inc(x_6); -lean_inc(x_258); -lean_inc(x_288); -x_290 = lean_simp(x_1, x_4, x_2, x_288, x_258, x_6, x_286, x_8, x_289); -if (lean_obj_tag(x_290) == 0) +lean_inc(x_153); +lean_inc(x_155); +x_157 = l_Lean_Meta_Simp_main_simpMain(x_1, x_4, x_2, x_155, x_153, x_6, x_7, x_8, x_156); +if (lean_obj_tag(x_157) == 0) { -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; uint8_t x_299; -lean_dec(x_273); -lean_dec(x_15); -x_291 = lean_ctor_get(x_290, 0); -lean_inc(x_291); -x_292 = lean_ctor_get(x_290, 1); -lean_inc(x_292); -lean_dec(x_290); -x_293 = lean_st_ref_get(x_288, x_292); -lean_dec(x_288); -x_294 = lean_ctor_get(x_293, 0); -lean_inc(x_294); -x_295 = lean_ctor_get(x_293, 1); -lean_inc(x_295); -lean_dec(x_293); -x_296 = l_Lean_Meta_Simp_main___closed__3; -x_297 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_296, x_258, x_6, x_286, x_8, x_295); -x_298 = lean_ctor_get(x_297, 0); -lean_inc(x_298); -x_299 = lean_unbox(x_298); -lean_dec(x_298); -if (x_299 == 0) +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; +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_st_ref_get(x_155, x_159); +lean_dec(x_155); +x_161 = lean_ctor_get(x_160, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_160, 1); +lean_inc(x_162); +lean_dec(x_160); +x_163 = l_Lean_Meta_Simp_main___closed__2; +x_164 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_163, x_153, x_6, x_7, x_8, x_162); +x_165 = lean_ctor_get(x_164, 0); +lean_inc(x_165); +x_166 = lean_unbox(x_165); +lean_dec(x_165); +if (x_166 == 0) { -lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; -x_300 = lean_ctor_get(x_297, 1); -lean_inc(x_300); -lean_dec(x_297); -x_301 = lean_box(0); -x_302 = l_Lean_Meta_Simp_main___lambda__1(x_294, x_291, x_301, x_258, x_6, x_286, x_8, x_300); +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +x_167 = lean_ctor_get(x_164, 1); +lean_inc(x_167); +lean_dec(x_164); +x_168 = lean_box(0); +x_169 = l_Lean_Meta_Simp_main___lambda__1(x_161, x_158, x_168, x_153, x_6, x_7, x_8, x_167); lean_dec(x_8); -lean_dec(x_286); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_258); -lean_dec(x_294); -x_303 = lean_ctor_get(x_302, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_302, 1); -lean_inc(x_304); -if (lean_is_exclusive(x_302)) { - lean_ctor_release(x_302, 0); - lean_ctor_release(x_302, 1); - x_305 = x_302; +lean_dec(x_153); +lean_dec(x_161); +x_170 = lean_ctor_get(x_169, 0); +lean_inc(x_170); +x_171 = lean_ctor_get(x_169, 1); +lean_inc(x_171); +if (lean_is_exclusive(x_169)) { + lean_ctor_release(x_169, 0); + lean_ctor_release(x_169, 1); + x_172 = x_169; } else { - lean_dec_ref(x_302); - x_305 = lean_box(0); + lean_dec_ref(x_169); + x_172 = lean_box(0); } -if (lean_is_scalar(x_305)) { - x_306 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_172)) { + x_173 = lean_alloc_ctor(0, 2, 0); } else { - x_306 = x_305; + x_173 = x_172; } -lean_ctor_set(x_306, 0, x_303); -lean_ctor_set(x_306, 1, x_304); -return x_306; +lean_ctor_set(x_173, 0, x_170); +lean_ctor_set(x_173, 1, x_171); +return x_173; } else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -x_307 = lean_ctor_get(x_297, 1); -lean_inc(x_307); -lean_dec(x_297); -x_308 = lean_ctor_get(x_294, 3); -lean_inc(x_308); -x_309 = l___private_Init_Data_Repr_0__Nat_reprFast(x_308); -x_310 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_310, 0, x_309); -x_311 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_311, 0, x_310); -x_312 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_313 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_313, 0, x_312); -lean_ctor_set(x_313, 1, x_311); -x_314 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_314, 0, x_313); -lean_ctor_set(x_314, 1, x_312); -x_315 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_296, x_314, x_258, x_6, x_286, x_8, x_307); -x_316 = lean_ctor_get(x_315, 0); -lean_inc(x_316); -x_317 = lean_ctor_get(x_315, 1); -lean_inc(x_317); -lean_dec(x_315); -x_318 = l_Lean_Meta_Simp_main___lambda__1(x_294, x_291, x_316, x_258, x_6, x_286, x_8, x_317); -lean_dec(x_8); -lean_dec(x_286); -lean_dec(x_6); -lean_dec(x_258); -lean_dec(x_316); -lean_dec(x_294); -x_319 = lean_ctor_get(x_318, 0); -lean_inc(x_319); -x_320 = lean_ctor_get(x_318, 1); -lean_inc(x_320); -if (lean_is_exclusive(x_318)) { - lean_ctor_release(x_318, 0); - lean_ctor_release(x_318, 1); - x_321 = x_318; +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_174 = lean_ctor_get(x_164, 1); +lean_inc(x_174); +lean_dec(x_164); +x_175 = lean_ctor_get(x_161, 3); +lean_inc(x_175); +x_176 = l___private_Init_Data_Repr_0__Nat_reprFast(x_175); +x_177 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_177, 0, x_176); +x_178 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_178, 0, x_177); +x_179 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +x_180 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_180, 0, x_179); +lean_ctor_set(x_180, 1, x_178); +x_181 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_179); +x_182 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_163, x_181, x_153, x_6, x_7, x_8, x_174); +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_ctor_get(x_182, 1); +lean_inc(x_184); +lean_dec(x_182); +x_185 = l_Lean_Meta_Simp_main___lambda__1(x_161, x_158, x_183, x_153, x_6, x_7, x_8, x_184); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_153); +lean_dec(x_183); +lean_dec(x_161); +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_188 = x_185; } else { - lean_dec_ref(x_318); - x_321 = lean_box(0); + lean_dec_ref(x_185); + x_188 = lean_box(0); } -if (lean_is_scalar(x_321)) { - x_322 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_188)) { + x_189 = lean_alloc_ctor(0, 2, 0); } else { - x_322 = x_321; + x_189 = x_188; } -lean_ctor_set(x_322, 0, x_319); -lean_ctor_set(x_322, 1, x_320); -return x_322; +lean_ctor_set(x_189, 0, x_186); +lean_ctor_set(x_189, 1, x_187); +return x_189; } } else { -lean_object* x_323; lean_object* x_324; -lean_dec(x_288); -lean_dec(x_286); -x_323 = lean_ctor_get(x_290, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_290, 1); -lean_inc(x_324); -lean_dec(x_290); -x_274 = x_323; -x_275 = x_324; -goto block_284; -} -block_284: -{ -uint8_t x_276; -x_276 = l_Lean_Exception_isRuntime(x_274); -if (x_276 == 0) -{ -lean_object* x_277; -lean_dec(x_273); -lean_dec(x_258); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_155); +lean_dec(x_153); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -if (lean_is_scalar(x_15)) { - x_277 = lean_alloc_ctor(1, 2, 0); -} else { - x_277 = x_15; - lean_ctor_set_tag(x_277, 1); -} -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_object* x_282; lean_object* x_283; -lean_dec(x_15); -x_278 = l_Lean_Meta_Simp_main___closed__1; -x_279 = l_Lean_Meta_throwNestedTacticEx___rarg(x_278, x_274, x_258, x_6, x_273, x_8, x_275); -x_280 = lean_ctor_get(x_279, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -if (lean_is_exclusive(x_279)) { - lean_ctor_release(x_279, 0); - lean_ctor_release(x_279, 1); - x_282 = x_279; +x_190 = lean_ctor_get(x_157, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_157, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_192 = x_157; } else { - lean_dec_ref(x_279); - x_282 = lean_box(0); + lean_dec_ref(x_157); + x_192 = lean_box(0); } -if (lean_is_scalar(x_282)) { - x_283 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); } else { - x_283 = x_282; -} -lean_ctor_set(x_283, 0, x_280); -lean_ctor_set(x_283, 1, x_281); -return x_283; + x_193 = x_192; } +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; } } } else { -lean_object* x_325; uint32_t x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; uint32_t x_330; lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; uint8_t x_346; uint8_t x_347; lean_object* x_348; uint8_t x_349; uint8_t x_350; uint8_t x_351; uint8_t x_352; uint8_t x_353; uint8_t x_354; uint8_t x_355; uint8_t x_356; uint8_t x_357; uint8_t x_358; uint8_t x_359; lean_object* x_360; uint8_t x_361; uint8_t x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; uint8_t x_376; lean_object* x_377; uint8_t x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; uint8_t x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; -x_325 = lean_ctor_get(x_2, 0); -x_326 = lean_ctor_get_uint32(x_2, sizeof(void*)*4); -x_327 = lean_ctor_get(x_2, 1); -x_328 = lean_ctor_get(x_2, 2); -x_329 = lean_ctor_get(x_2, 3); -x_330 = lean_ctor_get_uint32(x_2, sizeof(void*)*4 + 4); -lean_inc(x_329); -lean_inc(x_328); -lean_inc(x_327); -lean_inc(x_325); +lean_object* x_194; uint32_t x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; uint32_t 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; uint8_t x_209; uint8_t x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; uint8_t x_220; uint8_t x_221; uint8_t x_222; uint8_t x_223; uint8_t x_224; uint8_t x_225; uint8_t x_226; uint8_t x_227; uint8_t x_228; uint8_t x_229; uint8_t x_230; lean_object* x_231; uint8_t x_232; uint8_t 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_194 = lean_ctor_get(x_2, 0); +x_195 = lean_ctor_get_uint32(x_2, sizeof(void*)*5); +x_196 = lean_ctor_get(x_2, 1); +x_197 = lean_ctor_get(x_2, 2); +x_198 = lean_ctor_get(x_2, 3); +x_199 = lean_ctor_get_uint32(x_2, sizeof(void*)*5 + 4); +lean_inc(x_198); +lean_inc(x_197); +lean_inc(x_196); +lean_inc(x_194); lean_dec(x_2); -x_331 = l_Lean_Meta_Simp_Config_updateArith(x_325, x_7, x_8, x_9); -x_332 = lean_ctor_get(x_331, 0); -lean_inc(x_332); -x_333 = lean_ctor_get(x_331, 1); -lean_inc(x_333); -if (lean_is_exclusive(x_331)) { - lean_ctor_release(x_331, 0); - lean_ctor_release(x_331, 1); - x_334 = x_331; -} else { - lean_dec_ref(x_331); - x_334 = lean_box(0); -} -lean_inc(x_332); -x_335 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_335, 0, x_332); -lean_ctor_set(x_335, 1, x_327); -lean_ctor_set(x_335, 2, x_328); -lean_ctor_set(x_335, 3, x_329); -lean_ctor_set_uint32(x_335, sizeof(void*)*4, x_326); -lean_ctor_set_uint32(x_335, sizeof(void*)*4 + 4, x_330); -x_336 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; -x_337 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; -x_338 = lean_unsigned_to_nat(0u); -x_339 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_339, 0, x_336); -lean_ctor_set(x_339, 1, x_337); -lean_ctor_set(x_339, 2, x_3); -lean_ctor_set(x_339, 3, x_338); -x_340 = lean_ctor_get(x_5, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_5, 1); -lean_inc(x_341); -x_342 = lean_ctor_get(x_5, 2); -lean_inc(x_342); -x_343 = lean_ctor_get(x_5, 3); -lean_inc(x_343); -x_344 = lean_ctor_get(x_5, 4); -lean_inc(x_344); -x_345 = lean_ctor_get(x_5, 5); -lean_inc(x_345); -x_346 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); -x_347 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); +x_200 = l_Lean_Meta_Simp_Config_updateArith(x_194, x_7, x_8, x_9); +x_201 = lean_ctor_get(x_200, 0); +lean_inc(x_201); +x_202 = lean_ctor_get(x_200, 1); +lean_inc(x_202); +lean_dec(x_200); +x_203 = lean_ctor_get(x_5, 0); +lean_inc(x_203); +x_204 = lean_ctor_get(x_5, 1); +lean_inc(x_204); +x_205 = lean_ctor_get(x_5, 2); +lean_inc(x_205); +x_206 = lean_ctor_get(x_5, 3); +lean_inc(x_206); +x_207 = lean_ctor_get(x_5, 4); +lean_inc(x_207); +x_208 = lean_ctor_get(x_5, 5); +lean_inc(x_208); +x_209 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); +x_210 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); if (lean_is_exclusive(x_5)) { lean_ctor_release(x_5, 0); lean_ctor_release(x_5, 1); @@ -32715,333 +32923,236 @@ if (lean_is_exclusive(x_5)) { lean_ctor_release(x_5, 3); lean_ctor_release(x_5, 4); lean_ctor_release(x_5, 5); - x_348 = x_5; + x_211 = x_5; } else { lean_dec_ref(x_5); - x_348 = lean_box(0); -} -x_349 = lean_ctor_get_uint8(x_340, 0); -x_350 = lean_ctor_get_uint8(x_340, 1); -x_351 = lean_ctor_get_uint8(x_340, 2); -x_352 = lean_ctor_get_uint8(x_340, 3); -x_353 = lean_ctor_get_uint8(x_340, 4); -x_354 = lean_ctor_get_uint8(x_340, 5); -x_355 = lean_ctor_get_uint8(x_340, 6); -x_356 = lean_ctor_get_uint8(x_340, 7); -x_357 = lean_ctor_get_uint8(x_340, 8); -x_358 = lean_ctor_get_uint8(x_340, 10); -x_359 = lean_ctor_get_uint8(x_340, 12); -if (lean_is_exclusive(x_340)) { - x_360 = x_340; -} else { - lean_dec_ref(x_340); - x_360 = lean_box(0); -} -x_361 = lean_ctor_get_uint8(x_332, sizeof(void*)*2 + 6); -lean_dec(x_332); -x_362 = 2; -if (lean_is_scalar(x_360)) { - x_363 = lean_alloc_ctor(0, 0, 13); -} else { - x_363 = x_360; -} -lean_ctor_set_uint8(x_363, 0, x_349); -lean_ctor_set_uint8(x_363, 1, x_350); -lean_ctor_set_uint8(x_363, 2, x_351); -lean_ctor_set_uint8(x_363, 3, x_352); -lean_ctor_set_uint8(x_363, 4, x_353); -lean_ctor_set_uint8(x_363, 5, x_354); -lean_ctor_set_uint8(x_363, 6, x_355); -lean_ctor_set_uint8(x_363, 7, x_356); -lean_ctor_set_uint8(x_363, 8, x_357); -lean_ctor_set_uint8(x_363, 9, x_362); -lean_ctor_set_uint8(x_363, 10, x_358); -lean_ctor_set_uint8(x_363, 11, x_361); -lean_ctor_set_uint8(x_363, 12, x_359); -if (lean_is_scalar(x_348)) { - x_364 = lean_alloc_ctor(0, 6, 2); -} else { - x_364 = x_348; -} -lean_ctor_set(x_364, 0, x_363); -lean_ctor_set(x_364, 1, x_341); -lean_ctor_set(x_364, 2, x_342); -lean_ctor_set(x_364, 3, x_343); -lean_ctor_set(x_364, 4, x_344); -lean_ctor_set(x_364, 5, x_345); -lean_ctor_set_uint8(x_364, sizeof(void*)*6, x_346); -lean_ctor_set_uint8(x_364, sizeof(void*)*6 + 1, x_347); -x_365 = lean_ctor_get(x_7, 0); -lean_inc(x_365); -x_366 = lean_ctor_get(x_7, 1); -lean_inc(x_366); -x_367 = lean_ctor_get(x_7, 2); -lean_inc(x_367); -x_368 = lean_ctor_get(x_7, 3); -lean_inc(x_368); -x_369 = lean_ctor_get(x_7, 4); -lean_inc(x_369); -x_370 = lean_ctor_get(x_7, 5); -lean_inc(x_370); -x_371 = lean_ctor_get(x_7, 6); -lean_inc(x_371); -x_372 = lean_ctor_get(x_7, 7); -lean_inc(x_372); -x_373 = lean_ctor_get(x_7, 8); -lean_inc(x_373); -x_374 = lean_ctor_get(x_7, 9); -lean_inc(x_374); -x_375 = lean_ctor_get(x_7, 10); -lean_inc(x_375); -x_376 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - lean_ctor_release(x_7, 7); - lean_ctor_release(x_7, 8); - lean_ctor_release(x_7, 9); - lean_ctor_release(x_7, 10); - x_377 = x_7; + x_211 = lean_box(0); +} +lean_inc(x_204); +x_212 = lean_local_ctx_num_indices(x_204); +lean_inc(x_201); +x_213 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_213, 0, x_201); +lean_ctor_set(x_213, 1, x_196); +lean_ctor_set(x_213, 2, x_197); +lean_ctor_set(x_213, 3, x_198); +lean_ctor_set(x_213, 4, x_212); +lean_ctor_set_uint32(x_213, sizeof(void*)*5, x_195); +lean_ctor_set_uint32(x_213, sizeof(void*)*5 + 4, x_199); +x_214 = lean_ctor_get(x_3, 0); +x_215 = lean_ctor_get(x_3, 1); +x_216 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_217 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; +x_218 = lean_unsigned_to_nat(0u); +lean_inc(x_215); +lean_inc(x_214); +x_219 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_219, 0, x_216); +lean_ctor_set(x_219, 1, x_217); +lean_ctor_set(x_219, 2, x_214); +lean_ctor_set(x_219, 3, x_218); +lean_ctor_set(x_219, 4, x_215); +x_220 = lean_ctor_get_uint8(x_203, 0); +x_221 = lean_ctor_get_uint8(x_203, 1); +x_222 = lean_ctor_get_uint8(x_203, 2); +x_223 = lean_ctor_get_uint8(x_203, 3); +x_224 = lean_ctor_get_uint8(x_203, 4); +x_225 = lean_ctor_get_uint8(x_203, 5); +x_226 = lean_ctor_get_uint8(x_203, 6); +x_227 = lean_ctor_get_uint8(x_203, 7); +x_228 = lean_ctor_get_uint8(x_203, 8); +x_229 = lean_ctor_get_uint8(x_203, 10); +x_230 = lean_ctor_get_uint8(x_203, 12); +if (lean_is_exclusive(x_203)) { + x_231 = x_203; +} else { + lean_dec_ref(x_203); + x_231 = lean_box(0); +} +x_232 = lean_ctor_get_uint8(x_201, sizeof(void*)*2 + 6); +lean_dec(x_201); +x_233 = 2; +if (lean_is_scalar(x_231)) { + x_234 = lean_alloc_ctor(0, 0, 13); } else { - lean_dec_ref(x_7); - x_377 = lean_box(0); + x_234 = x_231; } -x_378 = 1; -lean_inc(x_375); -lean_inc(x_374); -lean_inc(x_373); -lean_inc(x_372); -lean_inc(x_371); -lean_inc(x_370); -lean_inc(x_369); -lean_inc(x_368); -lean_inc(x_367); -lean_inc(x_366); -lean_inc(x_365); -if (lean_is_scalar(x_377)) { - x_379 = lean_alloc_ctor(0, 11, 2); -} else { - x_379 = x_377; -} -lean_ctor_set(x_379, 0, x_365); -lean_ctor_set(x_379, 1, x_366); -lean_ctor_set(x_379, 2, x_367); -lean_ctor_set(x_379, 3, x_368); -lean_ctor_set(x_379, 4, x_369); -lean_ctor_set(x_379, 5, x_370); -lean_ctor_set(x_379, 6, x_371); -lean_ctor_set(x_379, 7, x_372); -lean_ctor_set(x_379, 8, x_373); -lean_ctor_set(x_379, 9, x_374); -lean_ctor_set(x_379, 10, x_375); -lean_ctor_set_uint8(x_379, sizeof(void*)*11, x_378); -lean_ctor_set_uint8(x_379, sizeof(void*)*11 + 1, x_376); -x_391 = 0; -x_392 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_392, 0, x_365); -lean_ctor_set(x_392, 1, x_366); -lean_ctor_set(x_392, 2, x_367); -lean_ctor_set(x_392, 3, x_368); -lean_ctor_set(x_392, 4, x_369); -lean_ctor_set(x_392, 5, x_370); -lean_ctor_set(x_392, 6, x_371); -lean_ctor_set(x_392, 7, x_372); -lean_ctor_set(x_392, 8, x_373); -lean_ctor_set(x_392, 9, x_374); -lean_ctor_set(x_392, 10, x_375); -lean_ctor_set_uint8(x_392, sizeof(void*)*11, x_391); -lean_ctor_set_uint8(x_392, sizeof(void*)*11 + 1, x_376); -x_393 = lean_st_mk_ref(x_339, x_333); -x_394 = lean_ctor_get(x_393, 0); -lean_inc(x_394); -x_395 = lean_ctor_get(x_393, 1); -lean_inc(x_395); -lean_dec(x_393); +lean_ctor_set_uint8(x_234, 0, x_220); +lean_ctor_set_uint8(x_234, 1, x_221); +lean_ctor_set_uint8(x_234, 2, x_222); +lean_ctor_set_uint8(x_234, 3, x_223); +lean_ctor_set_uint8(x_234, 4, x_224); +lean_ctor_set_uint8(x_234, 5, x_225); +lean_ctor_set_uint8(x_234, 6, x_226); +lean_ctor_set_uint8(x_234, 7, x_227); +lean_ctor_set_uint8(x_234, 8, x_228); +lean_ctor_set_uint8(x_234, 9, x_233); +lean_ctor_set_uint8(x_234, 10, x_229); +lean_ctor_set_uint8(x_234, 11, x_232); +lean_ctor_set_uint8(x_234, 12, x_230); +if (lean_is_scalar(x_211)) { + x_235 = lean_alloc_ctor(0, 6, 2); +} else { + x_235 = x_211; +} +lean_ctor_set(x_235, 0, x_234); +lean_ctor_set(x_235, 1, x_204); +lean_ctor_set(x_235, 2, x_205); +lean_ctor_set(x_235, 3, x_206); +lean_ctor_set(x_235, 4, x_207); +lean_ctor_set(x_235, 5, x_208); +lean_ctor_set_uint8(x_235, sizeof(void*)*6, x_209); +lean_ctor_set_uint8(x_235, sizeof(void*)*6 + 1, x_210); +x_236 = lean_st_mk_ref(x_219, x_202); +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_8); -lean_inc(x_392); +lean_inc(x_7); lean_inc(x_6); -lean_inc(x_364); -lean_inc(x_394); -x_396 = lean_simp(x_1, x_4, x_335, x_394, x_364, x_6, x_392, x_8, x_395); -if (lean_obj_tag(x_396) == 0) +lean_inc(x_235); +lean_inc(x_237); +x_239 = l_Lean_Meta_Simp_main_simpMain(x_1, x_4, x_213, x_237, x_235, x_6, x_7, x_8, x_238); +if (lean_obj_tag(x_239) == 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; lean_object* x_403; lean_object* x_404; uint8_t x_405; -lean_dec(x_379); -lean_dec(x_334); -x_397 = lean_ctor_get(x_396, 0); -lean_inc(x_397); -x_398 = lean_ctor_get(x_396, 1); -lean_inc(x_398); -lean_dec(x_396); -x_399 = lean_st_ref_get(x_394, x_398); -lean_dec(x_394); -x_400 = lean_ctor_get(x_399, 0); -lean_inc(x_400); -x_401 = lean_ctor_get(x_399, 1); -lean_inc(x_401); -lean_dec(x_399); -x_402 = l_Lean_Meta_Simp_main___closed__3; -x_403 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_402, x_364, x_6, x_392, x_8, x_401); -x_404 = lean_ctor_get(x_403, 0); -lean_inc(x_404); -x_405 = lean_unbox(x_404); -lean_dec(x_404); -if (x_405 == 0) -{ -lean_object* x_406; lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; -x_406 = lean_ctor_get(x_403, 1); -lean_inc(x_406); -lean_dec(x_403); -x_407 = lean_box(0); -x_408 = l_Lean_Meta_Simp_main___lambda__1(x_400, x_397, x_407, x_364, x_6, x_392, x_8, x_406); -lean_dec(x_8); -lean_dec(x_392); -lean_dec(x_6); -lean_dec(x_364); -lean_dec(x_400); -x_409 = lean_ctor_get(x_408, 0); -lean_inc(x_409); -x_410 = lean_ctor_get(x_408, 1); -lean_inc(x_410); -if (lean_is_exclusive(x_408)) { - lean_ctor_release(x_408, 0); - lean_ctor_release(x_408, 1); - x_411 = x_408; -} else { - lean_dec_ref(x_408); - x_411 = lean_box(0); -} -if (lean_is_scalar(x_411)) { - x_412 = lean_alloc_ctor(0, 2, 0); -} else { - x_412 = x_411; -} -lean_ctor_set(x_412, 0, x_409); -lean_ctor_set(x_412, 1, x_410); -return x_412; -} -else +lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; +x_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 = lean_st_ref_get(x_237, x_241); +lean_dec(x_237); +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); +lean_inc(x_244); +lean_dec(x_242); +x_245 = l_Lean_Meta_Simp_main___closed__2; +x_246 = l_Lean_isTracingEnabledFor___at_Lean_Meta_processPostponed_loop___spec__1(x_245, x_235, x_6, x_7, x_8, x_244); +x_247 = lean_ctor_get(x_246, 0); +lean_inc(x_247); +x_248 = lean_unbox(x_247); +lean_dec(x_247); +if (x_248 == 0) { -lean_object* x_413; lean_object* x_414; lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; lean_object* x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; -x_413 = lean_ctor_get(x_403, 1); -lean_inc(x_413); -lean_dec(x_403); -x_414 = lean_ctor_get(x_400, 3); -lean_inc(x_414); -x_415 = l___private_Init_Data_Repr_0__Nat_reprFast(x_414); -x_416 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_416, 0, x_415); -x_417 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_417, 0, x_416); -x_418 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; -x_419 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_419, 0, x_418); -lean_ctor_set(x_419, 1, x_417); -x_420 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_420, 0, x_419); -lean_ctor_set(x_420, 1, x_418); -x_421 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_402, x_420, x_364, x_6, x_392, x_8, x_413); -x_422 = lean_ctor_get(x_421, 0); -lean_inc(x_422); -x_423 = lean_ctor_get(x_421, 1); -lean_inc(x_423); -lean_dec(x_421); -x_424 = l_Lean_Meta_Simp_main___lambda__1(x_400, x_397, x_422, x_364, x_6, x_392, x_8, x_423); +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_ctor_get(x_246, 1); +lean_inc(x_249); +lean_dec(x_246); +x_250 = lean_box(0); +x_251 = l_Lean_Meta_Simp_main___lambda__1(x_243, x_240, x_250, x_235, x_6, x_7, x_8, x_249); lean_dec(x_8); -lean_dec(x_392); +lean_dec(x_7); lean_dec(x_6); -lean_dec(x_364); -lean_dec(x_422); -lean_dec(x_400); -x_425 = lean_ctor_get(x_424, 0); -lean_inc(x_425); -x_426 = lean_ctor_get(x_424, 1); -lean_inc(x_426); -if (lean_is_exclusive(x_424)) { - lean_ctor_release(x_424, 0); - lean_ctor_release(x_424, 1); - x_427 = x_424; +lean_dec(x_235); +lean_dec(x_243); +x_252 = lean_ctor_get(x_251, 0); +lean_inc(x_252); +x_253 = lean_ctor_get(x_251, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_251)) { + lean_ctor_release(x_251, 0); + lean_ctor_release(x_251, 1); + x_254 = x_251; } else { - lean_dec_ref(x_424); - x_427 = lean_box(0); + lean_dec_ref(x_251); + x_254 = lean_box(0); } -if (lean_is_scalar(x_427)) { - x_428 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_254)) { + x_255 = lean_alloc_ctor(0, 2, 0); } else { - x_428 = x_427; -} -lean_ctor_set(x_428, 0, x_425); -lean_ctor_set(x_428, 1, x_426); -return x_428; + x_255 = x_254; } +lean_ctor_set(x_255, 0, x_252); +lean_ctor_set(x_255, 1, x_253); +return x_255; } else { -lean_object* x_429; lean_object* x_430; -lean_dec(x_394); -lean_dec(x_392); -x_429 = lean_ctor_get(x_396, 0); -lean_inc(x_429); -x_430 = lean_ctor_get(x_396, 1); -lean_inc(x_430); -lean_dec(x_396); -x_380 = x_429; -x_381 = x_430; -goto block_390; -} -block_390: -{ -uint8_t x_382; -x_382 = l_Lean_Exception_isRuntime(x_380); -if (x_382 == 0) -{ -lean_object* x_383; -lean_dec(x_379); -lean_dec(x_364); +lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +x_256 = lean_ctor_get(x_246, 1); +lean_inc(x_256); +lean_dec(x_246); +x_257 = lean_ctor_get(x_243, 3); +lean_inc(x_257); +x_258 = l___private_Init_Data_Repr_0__Nat_reprFast(x_257); +x_259 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_259, 0, x_258); +x_260 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_260, 0, x_259); +x_261 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__13; +x_262 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_262, 0, x_261); +lean_ctor_set(x_262, 1, x_260); +x_263 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_263, 0, x_262); +lean_ctor_set(x_263, 1, x_261); +x_264 = l_Lean_addTrace___at_Lean_Meta_processPostponed_loop___spec__2(x_245, x_263, x_235, x_6, x_7, x_8, x_256); +x_265 = lean_ctor_get(x_264, 0); +lean_inc(x_265); +x_266 = lean_ctor_get(x_264, 1); +lean_inc(x_266); +lean_dec(x_264); +x_267 = l_Lean_Meta_Simp_main___lambda__1(x_243, x_240, x_265, x_235, x_6, x_7, x_8, x_266); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -if (lean_is_scalar(x_334)) { - x_383 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_235); +lean_dec(x_265); +lean_dec(x_243); +x_268 = lean_ctor_get(x_267, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_267, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_267)) { + lean_ctor_release(x_267, 0); + lean_ctor_release(x_267, 1); + x_270 = x_267; +} else { + lean_dec_ref(x_267); + x_270 = lean_box(0); +} +if (lean_is_scalar(x_270)) { + x_271 = lean_alloc_ctor(0, 2, 0); } else { - x_383 = x_334; - lean_ctor_set_tag(x_383, 1); + x_271 = x_270; +} +lean_ctor_set(x_271, 0, x_268); +lean_ctor_set(x_271, 1, x_269); +return x_271; } -lean_ctor_set(x_383, 0, x_380); -lean_ctor_set(x_383, 1, x_381); -return x_383; } else { -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_dec(x_334); -x_384 = l_Lean_Meta_Simp_main___closed__1; -x_385 = l_Lean_Meta_throwNestedTacticEx___rarg(x_384, x_380, x_364, x_6, x_379, x_8, x_381); -x_386 = lean_ctor_get(x_385, 0); -lean_inc(x_386); -x_387 = lean_ctor_get(x_385, 1); -lean_inc(x_387); -if (lean_is_exclusive(x_385)) { - lean_ctor_release(x_385, 0); - lean_ctor_release(x_385, 1); - x_388 = x_385; +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; +lean_dec(x_237); +lean_dec(x_235); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_272 = lean_ctor_get(x_239, 0); +lean_inc(x_272); +x_273 = lean_ctor_get(x_239, 1); +lean_inc(x_273); +if (lean_is_exclusive(x_239)) { + lean_ctor_release(x_239, 0); + lean_ctor_release(x_239, 1); + x_274 = x_239; } else { - lean_dec_ref(x_385); - x_388 = lean_box(0); + lean_dec_ref(x_239); + x_274 = lean_box(0); } -if (lean_is_scalar(x_388)) { - x_389 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_274)) { + x_275 = lean_alloc_ctor(1, 2, 0); } else { - x_389 = x_388; -} -lean_ctor_set(x_389, 0, x_386); -lean_ctor_set(x_389, 1, x_387); -return x_389; + x_275 = x_274; } +lean_ctor_set(x_275, 0, x_272); +lean_ctor_set(x_275, 1, x_273); +return x_275; } } } @@ -33060,343 +33171,403 @@ lean_dec(x_1); return x_9; } } -static lean_object* _init_l_Lean_Meta_Simp_dsimpMain___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("dsimp", 5); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_Simp_dsimpMain___closed__2() { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_main___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_Meta_Simp_dsimpMain___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_10; +x_10 = l_Lean_Meta_Simp_main(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimpMain(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_Simp_dsimpMain_dsimpMain(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; -x_11 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_13, 0, x_10); -lean_ctor_set(x_13, 1, x_11); -lean_ctor_set(x_13, 2, x_3); -lean_ctor_set(x_13, 3, x_12); -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, 0); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; -x_17 = lean_ctor_get(x_2, 0); -lean_inc(x_17); -x_18 = lean_ctor_get_uint8(x_17, sizeof(void*)*2 + 6); -lean_dec(x_17); -x_19 = 2; -lean_ctor_set_uint8(x_15, 9, x_19); -lean_ctor_set_uint8(x_15, 11, x_18); -x_20 = !lean_is_exclusive(x_7); -if (x_20 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 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; uint8_t x_32; uint8_t x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_21 = lean_ctor_get(x_7, 0); -x_22 = lean_ctor_get(x_7, 1); -x_23 = lean_ctor_get(x_7, 2); -x_24 = lean_ctor_get(x_7, 3); -x_25 = lean_ctor_get(x_7, 4); -x_26 = lean_ctor_get(x_7, 5); -x_27 = lean_ctor_get(x_7, 6); -x_28 = lean_ctor_get(x_7, 7); -x_29 = lean_ctor_get(x_7, 8); -x_30 = lean_ctor_get(x_7, 9); -x_31 = lean_ctor_get(x_7, 10); -x_32 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -x_33 = 1; -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_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; uint8_t x_23; uint8_t x_24; lean_object* x_25; lean_object* x_26; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +x_13 = lean_ctor_get(x_7, 2); +x_14 = lean_ctor_get(x_7, 3); +x_15 = lean_ctor_get(x_7, 4); +x_16 = lean_ctor_get(x_7, 5); +x_17 = lean_ctor_get(x_7, 6); +x_18 = lean_ctor_get(x_7, 7); +x_19 = lean_ctor_get(x_7, 8); +x_20 = lean_ctor_get(x_7, 9); +x_21 = lean_ctor_get(x_7, 10); +x_22 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); +x_23 = 1; lean_inc(x_21); -lean_ctor_set_uint8(x_7, sizeof(void*)*11, x_33); -x_34 = 0; -x_35 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_35, 0, x_21); -lean_ctor_set(x_35, 1, x_22); -lean_ctor_set(x_35, 2, x_23); -lean_ctor_set(x_35, 3, x_24); -lean_ctor_set(x_35, 4, x_25); -lean_ctor_set(x_35, 5, x_26); -lean_ctor_set(x_35, 6, x_27); -lean_ctor_set(x_35, 7, x_28); -lean_ctor_set(x_35, 8, x_29); -lean_ctor_set(x_35, 9, x_30); -lean_ctor_set(x_35, 10, x_31); -lean_ctor_set_uint8(x_35, sizeof(void*)*11, x_34); -lean_ctor_set_uint8(x_35, sizeof(void*)*11 + 1, x_32); -x_36 = lean_st_mk_ref(x_13, x_9); -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_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_12); +lean_inc(x_11); +lean_ctor_set_uint8(x_7, sizeof(void*)*11, x_23); +x_24 = 0; +x_25 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_25, 0, x_11); +lean_ctor_set(x_25, 1, x_12); +lean_ctor_set(x_25, 2, x_13); +lean_ctor_set(x_25, 3, x_14); +lean_ctor_set(x_25, 4, x_15); +lean_ctor_set(x_25, 5, x_16); +lean_ctor_set(x_25, 6, x_17); +lean_ctor_set(x_25, 7, x_18); +lean_ctor_set(x_25, 8, x_19); +lean_ctor_set(x_25, 9, x_20); +lean_ctor_set(x_25, 10, x_21); +lean_ctor_set_uint8(x_25, sizeof(void*)*11, x_24); +lean_ctor_set_uint8(x_25, sizeof(void*)*11 + 1, x_22); lean_inc(x_8); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_37); -x_39 = lean_dsimp(x_1, x_4, x_2, x_37, x_5, x_6, x_35, x_8, x_38); -if (lean_obj_tag(x_39) == 0) +lean_inc(x_4); +x_26 = lean_dsimp(x_1, x_2, x_3, x_4, x_5, x_6, x_25, x_8, x_9); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; +uint8_t x_27; lean_dec(x_7); -lean_dec(x_5); lean_dec(x_8); lean_dec(x_6); -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_st_ref_get(x_37, x_41); -lean_dec(x_37); -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 0) +lean_dec(x_5); +lean_dec(x_4); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) { -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_44, 2); -lean_inc(x_45); -lean_dec(x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_40); -lean_ctor_set(x_46, 1, x_45); -lean_ctor_set(x_42, 0, x_46); -return x_42; +return x_26; } 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 = lean_ctor_get(x_47, 2); -lean_inc(x_49); -lean_dec(x_47); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_40); -lean_ctor_set(x_50, 1, x_49); -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_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; } } else { -uint8_t x_52; -lean_dec(x_37); -x_52 = !lean_is_exclusive(x_39); -if (x_52 == 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; +x_31 = lean_ctor_get(x_26, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_26, 1); +lean_inc(x_32); +lean_dec(x_26); +x_33 = lean_st_ref_get(x_4, x_32); +lean_dec(x_4); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_34, 4); +lean_inc(x_36); +lean_dec(x_34); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_37 = l_Lean_Meta_Simp_reportDiag(x_36, x_5, x_6, x_7, x_8, x_35); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_53; lean_object* x_54; uint8_t x_55; -x_53 = lean_ctor_get(x_39, 0); -x_54 = lean_ctor_get(x_39, 1); -x_55 = l_Lean_Exception_isRuntime(x_53); -if (x_55 == 0) +uint8_t x_38; +x_38 = !lean_is_exclusive(x_37); +if (x_38 == 0) +{ +lean_object* x_39; lean_object* x_40; uint8_t x_41; +x_39 = lean_ctor_get(x_37, 1); +x_40 = lean_ctor_get(x_37, 0); +lean_dec(x_40); +x_41 = l_Lean_Exception_isRuntime(x_31); +if (x_41 == 0) { lean_dec(x_7); -lean_dec(x_5); lean_dec(x_8); lean_dec(x_6); -return x_39; +lean_dec(x_5); +lean_ctor_set_tag(x_37, 1); +lean_ctor_set(x_37, 0, x_31); +return x_37; } else { -lean_object* x_56; lean_object* x_57; uint8_t x_58; -lean_free_object(x_39); -x_56 = l_Lean_Meta_Simp_dsimpMain___closed__2; -x_57 = l_Lean_Meta_throwNestedTacticEx___rarg(x_56, x_53, x_5, x_6, x_7, x_8, x_54); -x_58 = !lean_is_exclusive(x_57); -if (x_58 == 0) +lean_object* x_42; lean_object* x_43; uint8_t x_44; +lean_free_object(x_37); +x_42 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_43 = l_Lean_Meta_throwNestedTacticEx___rarg(x_42, x_31, x_5, x_6, x_7, x_8, x_39); +x_44 = !lean_is_exclusive(x_43); +if (x_44 == 0) { -return x_57; +return x_43; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_57, 0); -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_57); -x_61 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_61, 0, x_59); -lean_ctor_set(x_61, 1, x_60); -return x_61; +lean_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_62; lean_object* x_63; uint8_t x_64; -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); -x_64 = l_Lean_Exception_isRuntime(x_62); -if (x_64 == 0) +lean_object* x_48; uint8_t x_49; +x_48 = lean_ctor_get(x_37, 1); +lean_inc(x_48); +lean_dec(x_37); +x_49 = l_Lean_Exception_isRuntime(x_31); +if (x_49 == 0) { -lean_object* x_65; +lean_object* x_50; lean_dec(x_7); +lean_dec(x_8); +lean_dec(x_6); lean_dec(x_5); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_31); +lean_ctor_set(x_50, 1, x_48); +return x_50; +} +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_51 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_52 = l_Lean_Meta_throwNestedTacticEx___rarg(x_51, x_31, x_5, x_6, x_7, x_8, x_48); +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_52); + x_55 = lean_box(0); +} +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(1, 2, 0); +} else { + x_56 = x_55; +} +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; +} +} +} +else +{ +uint8_t x_57; +lean_dec(x_31); +lean_dec(x_7); lean_dec(x_8); lean_dec(x_6); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_62); -lean_ctor_set(x_65, 1, x_63); -return x_65; +lean_dec(x_5); +x_57 = !lean_is_exclusive(x_37); +if (x_57 == 0) +{ +return x_37; } else { -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_66 = l_Lean_Meta_Simp_dsimpMain___closed__2; -x_67 = l_Lean_Meta_throwNestedTacticEx___rarg(x_66, x_62, x_5, x_6, x_7, x_8, x_63); -x_68 = lean_ctor_get(x_67, 0); +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_37, 0); +x_59 = lean_ctor_get(x_37, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_37); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +} +else +{ +lean_object* x_61; 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; uint8_t x_72; uint8_t x_73; lean_object* x_74; uint8_t x_75; lean_object* x_76; lean_object* x_77; +x_61 = lean_ctor_get(x_7, 0); +x_62 = lean_ctor_get(x_7, 1); +x_63 = lean_ctor_get(x_7, 2); +x_64 = lean_ctor_get(x_7, 3); +x_65 = lean_ctor_get(x_7, 4); +x_66 = lean_ctor_get(x_7, 5); +x_67 = lean_ctor_get(x_7, 6); +x_68 = lean_ctor_get(x_7, 7); +x_69 = lean_ctor_get(x_7, 8); +x_70 = lean_ctor_get(x_7, 9); +x_71 = lean_ctor_get(x_7, 10); +x_72 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); +lean_inc(x_71); +lean_inc(x_70); +lean_inc(x_69); lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); +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_dec(x_7); +x_73 = 1; +lean_inc(x_71); +lean_inc(x_70); lean_inc(x_69); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_70 = x_67; +lean_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); +x_74 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_74, 0, x_61); +lean_ctor_set(x_74, 1, x_62); +lean_ctor_set(x_74, 2, x_63); +lean_ctor_set(x_74, 3, x_64); +lean_ctor_set(x_74, 4, x_65); +lean_ctor_set(x_74, 5, x_66); +lean_ctor_set(x_74, 6, x_67); +lean_ctor_set(x_74, 7, x_68); +lean_ctor_set(x_74, 8, x_69); +lean_ctor_set(x_74, 9, x_70); +lean_ctor_set(x_74, 10, x_71); +lean_ctor_set_uint8(x_74, sizeof(void*)*11, x_73); +lean_ctor_set_uint8(x_74, sizeof(void*)*11 + 1, x_72); +x_75 = 0; +x_76 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_76, 0, x_61); +lean_ctor_set(x_76, 1, x_62); +lean_ctor_set(x_76, 2, x_63); +lean_ctor_set(x_76, 3, x_64); +lean_ctor_set(x_76, 4, x_65); +lean_ctor_set(x_76, 5, x_66); +lean_ctor_set(x_76, 6, x_67); +lean_ctor_set(x_76, 7, x_68); +lean_ctor_set(x_76, 8, x_69); +lean_ctor_set(x_76, 9, x_70); +lean_ctor_set(x_76, 10, x_71); +lean_ctor_set_uint8(x_76, sizeof(void*)*11, x_75); +lean_ctor_set_uint8(x_76, sizeof(void*)*11 + 1, x_72); +lean_inc(x_8); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_77 = lean_dsimp(x_1, x_2, x_3, x_4, x_5, x_6, x_76, x_8, x_9); +if (lean_obj_tag(x_77) == 0) +{ +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_74); +lean_dec(x_8); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_78 = lean_ctor_get(x_77, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_77, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_77)) { + lean_ctor_release(x_77, 0); + lean_ctor_release(x_77, 1); + x_80 = x_77; } else { - lean_dec_ref(x_67); - x_70 = lean_box(0); + lean_dec_ref(x_77); + x_80 = lean_box(0); } -if (lean_is_scalar(x_70)) { - x_71 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(0, 2, 0); } else { - x_71 = x_70; -} -lean_ctor_set(x_71, 0, x_68); -lean_ctor_set(x_71, 1, x_69); -return x_71; -} -} + x_81 = x_80; } +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; uint8_t x_83; uint8_t x_84; lean_object* x_85; uint8_t x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; -x_72 = lean_ctor_get(x_7, 0); -x_73 = lean_ctor_get(x_7, 1); -x_74 = lean_ctor_get(x_7, 2); -x_75 = lean_ctor_get(x_7, 3); -x_76 = lean_ctor_get(x_7, 4); -x_77 = lean_ctor_get(x_7, 5); -x_78 = lean_ctor_get(x_7, 6); -x_79 = lean_ctor_get(x_7, 7); -x_80 = lean_ctor_get(x_7, 8); -x_81 = lean_ctor_get(x_7, 9); -x_82 = lean_ctor_get(x_7, 10); -x_83 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_7); -x_84 = 1; +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_82 = lean_ctor_get(x_77, 0); lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_inc(x_79); -lean_inc(x_78); -lean_inc(x_77); -lean_inc(x_76); -lean_inc(x_75); -lean_inc(x_74); -lean_inc(x_73); -lean_inc(x_72); -x_85 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_85, 0, x_72); -lean_ctor_set(x_85, 1, x_73); -lean_ctor_set(x_85, 2, x_74); -lean_ctor_set(x_85, 3, x_75); -lean_ctor_set(x_85, 4, x_76); -lean_ctor_set(x_85, 5, x_77); -lean_ctor_set(x_85, 6, x_78); -lean_ctor_set(x_85, 7, x_79); -lean_ctor_set(x_85, 8, x_80); -lean_ctor_set(x_85, 9, x_81); -lean_ctor_set(x_85, 10, x_82); -lean_ctor_set_uint8(x_85, sizeof(void*)*11, x_84); -lean_ctor_set_uint8(x_85, sizeof(void*)*11 + 1, x_83); -x_86 = 0; -x_87 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_87, 0, x_72); -lean_ctor_set(x_87, 1, x_73); -lean_ctor_set(x_87, 2, x_74); -lean_ctor_set(x_87, 3, x_75); -lean_ctor_set(x_87, 4, x_76); -lean_ctor_set(x_87, 5, x_77); -lean_ctor_set(x_87, 6, x_78); -lean_ctor_set(x_87, 7, x_79); -lean_ctor_set(x_87, 8, x_80); -lean_ctor_set(x_87, 9, x_81); -lean_ctor_set(x_87, 10, x_82); -lean_ctor_set_uint8(x_87, sizeof(void*)*11, x_86); -lean_ctor_set_uint8(x_87, sizeof(void*)*11 + 1, x_83); -x_88 = lean_st_mk_ref(x_13, x_9); -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_83 = lean_ctor_get(x_77, 1); +lean_inc(x_83); +lean_dec(x_77); +x_84 = lean_st_ref_get(x_4, x_83); +lean_dec(x_4); +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, 4); +lean_inc(x_87); +lean_dec(x_85); lean_inc(x_8); +lean_inc(x_74); lean_inc(x_6); lean_inc(x_5); +x_88 = l_Lean_Meta_Simp_reportDiag(x_87, x_5, x_6, x_74, x_8, x_86); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; uint8_t x_91; +x_89 = lean_ctor_get(x_88, 1); lean_inc(x_89); -x_91 = lean_dsimp(x_1, x_4, x_2, x_89, x_5, x_6, x_87, x_8, x_90); -if (lean_obj_tag(x_91) == 0) +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_90 = x_88; +} else { + lean_dec_ref(x_88); + x_90 = lean_box(0); +} +x_91 = l_Lean_Exception_isRuntime(x_82); +if (x_91 == 0) { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -lean_dec(x_85); -lean_dec(x_5); +lean_object* x_92; +lean_dec(x_74); lean_dec(x_8); lean_dec(x_6); -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_st_ref_get(x_89, x_93); -lean_dec(x_89); +lean_dec(x_5); +if (lean_is_scalar(x_90)) { + x_92 = lean_alloc_ctor(1, 2, 0); +} else { + x_92 = x_90; + lean_ctor_set_tag(x_92, 1); +} +lean_ctor_set(x_92, 0, x_82); +lean_ctor_set(x_92, 1, x_89); +return x_92; +} +else +{ +lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_dec(x_90); +x_93 = l_Lean_Meta_Simp_main_simpMain___closed__1; +x_94 = l_Lean_Meta_throwNestedTacticEx___rarg(x_93, x_82, x_5, x_6, x_74, x_8, x_89); x_95 = lean_ctor_get(x_94, 0); lean_inc(x_95); x_96 = lean_ctor_get(x_94, 1); @@ -33409,585 +33580,432 @@ if (lean_is_exclusive(x_94)) { lean_dec_ref(x_94); x_97 = lean_box(0); } -x_98 = lean_ctor_get(x_95, 2); -lean_inc(x_98); -lean_dec(x_95); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_92); -lean_ctor_set(x_99, 1, x_98); if (lean_is_scalar(x_97)) { - x_100 = lean_alloc_ctor(0, 2, 0); + x_98 = lean_alloc_ctor(1, 2, 0); } else { - x_100 = x_97; + x_98 = x_97; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_96); -return x_100; +lean_ctor_set(x_98, 0, x_95); +lean_ctor_set(x_98, 1, x_96); +return x_98; } -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; -lean_dec(x_89); -x_101 = lean_ctor_get(x_91, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_91, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_103 = x_91; -} else { - lean_dec_ref(x_91); - x_103 = lean_box(0); } -x_104 = l_Lean_Exception_isRuntime(x_101); -if (x_104 == 0) +else { -lean_object* x_105; -lean_dec(x_85); -lean_dec(x_5); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; +lean_dec(x_82); +lean_dec(x_74); lean_dec(x_8); lean_dec(x_6); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_5); +x_99 = lean_ctor_get(x_88, 0); +lean_inc(x_99); +x_100 = lean_ctor_get(x_88, 1); +lean_inc(x_100); +if (lean_is_exclusive(x_88)) { + lean_ctor_release(x_88, 0); + lean_ctor_release(x_88, 1); + x_101 = x_88; } else { - x_105 = x_103; -} -lean_ctor_set(x_105, 0, x_101); -lean_ctor_set(x_105, 1, x_102); -return x_105; + lean_dec_ref(x_88); + x_101 = lean_box(0); } -else -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -lean_dec(x_103); -x_106 = l_Lean_Meta_Simp_dsimpMain___closed__2; -x_107 = l_Lean_Meta_throwNestedTacticEx___rarg(x_106, x_101, x_5, x_6, x_85, x_8, x_102); -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - x_110 = x_107; +if (lean_is_scalar(x_101)) { + x_102 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_107); - x_110 = lean_box(0); + x_102 = x_101; } -if (lean_is_scalar(x_110)) { - x_111 = lean_alloc_ctor(1, 2, 0); -} else { - x_111 = x_110; +lean_ctor_set(x_102, 0, x_99); +lean_ctor_set(x_102, 1, x_100); +return x_102; } -lean_ctor_set(x_111, 0, x_108); -lean_ctor_set(x_111, 1, x_109); -return x_111; } } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimpMain(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_10 = lean_ctor_get(x_3, 0); +x_11 = lean_ctor_get(x_3, 1); +x_12 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__5; +x_13 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__1; +x_14 = lean_unsigned_to_nat(0u); +lean_inc(x_11); +lean_inc(x_10); +x_15 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_15, 0, x_12); +lean_ctor_set(x_15, 1, x_13); +lean_ctor_set(x_15, 2, x_10); +lean_ctor_set(x_15, 3, x_14); +lean_ctor_set(x_15, 4, x_11); +x_16 = !lean_is_exclusive(x_5); +if (x_16 == 0) +{ +lean_object* x_17; uint8_t x_18; +x_17 = lean_ctor_get(x_5, 0); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_19 = lean_ctor_get(x_2, 0); +lean_inc(x_19); +x_20 = lean_ctor_get_uint8(x_19, sizeof(void*)*2 + 6); +lean_dec(x_19); +x_21 = 2; +lean_ctor_set_uint8(x_17, 9, x_21); +lean_ctor_set_uint8(x_17, 11, x_20); +x_22 = lean_st_mk_ref(x_15, x_9); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +lean_inc(x_23); +x_25 = l_Lean_Meta_Simp_dsimpMain_dsimpMain(x_1, x_4, x_2, x_23, x_5, x_6, x_7, x_8, x_24); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_25, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = lean_st_ref_get(x_23, x_27); +lean_dec(x_23); +x_29 = !lean_is_exclusive(x_28); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_30 = lean_ctor_get(x_28, 0); +x_31 = lean_ctor_get(x_30, 2); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 4); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_26); +lean_ctor_set(x_34, 1, x_33); +lean_ctor_set(x_28, 0, x_34); +return x_28; +} else { -uint8_t x_112; uint8_t x_113; uint8_t x_114; uint8_t x_115; uint8_t x_116; uint8_t x_117; uint8_t x_118; uint8_t x_119; uint8_t x_120; uint8_t x_121; uint8_t x_122; lean_object* x_123; uint8_t x_124; uint8_t 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; uint8_t x_138; lean_object* x_139; uint8_t x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_112 = lean_ctor_get_uint8(x_15, 0); -x_113 = lean_ctor_get_uint8(x_15, 1); -x_114 = lean_ctor_get_uint8(x_15, 2); -x_115 = lean_ctor_get_uint8(x_15, 3); -x_116 = lean_ctor_get_uint8(x_15, 4); -x_117 = lean_ctor_get_uint8(x_15, 5); -x_118 = lean_ctor_get_uint8(x_15, 6); -x_119 = lean_ctor_get_uint8(x_15, 7); -x_120 = lean_ctor_get_uint8(x_15, 8); -x_121 = lean_ctor_get_uint8(x_15, 10); -x_122 = lean_ctor_get_uint8(x_15, 12); -lean_dec(x_15); -x_123 = lean_ctor_get(x_2, 0); -lean_inc(x_123); -x_124 = lean_ctor_get_uint8(x_123, sizeof(void*)*2 + 6); -lean_dec(x_123); -x_125 = 2; -x_126 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_126, 0, x_112); -lean_ctor_set_uint8(x_126, 1, x_113); -lean_ctor_set_uint8(x_126, 2, x_114); -lean_ctor_set_uint8(x_126, 3, x_115); -lean_ctor_set_uint8(x_126, 4, x_116); -lean_ctor_set_uint8(x_126, 5, x_117); -lean_ctor_set_uint8(x_126, 6, x_118); -lean_ctor_set_uint8(x_126, 7, x_119); -lean_ctor_set_uint8(x_126, 8, x_120); -lean_ctor_set_uint8(x_126, 9, x_125); -lean_ctor_set_uint8(x_126, 10, x_121); -lean_ctor_set_uint8(x_126, 11, x_124); -lean_ctor_set_uint8(x_126, 12, x_122); -lean_ctor_set(x_5, 0, x_126); -x_127 = lean_ctor_get(x_7, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_7, 1); -lean_inc(x_128); -x_129 = lean_ctor_get(x_7, 2); -lean_inc(x_129); -x_130 = lean_ctor_get(x_7, 3); -lean_inc(x_130); -x_131 = lean_ctor_get(x_7, 4); -lean_inc(x_131); -x_132 = lean_ctor_get(x_7, 5); -lean_inc(x_132); -x_133 = lean_ctor_get(x_7, 6); -lean_inc(x_133); -x_134 = lean_ctor_get(x_7, 7); -lean_inc(x_134); -x_135 = lean_ctor_get(x_7, 8); -lean_inc(x_135); -x_136 = lean_ctor_get(x_7, 9); -lean_inc(x_136); -x_137 = lean_ctor_get(x_7, 10); -lean_inc(x_137); -x_138 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - lean_ctor_release(x_7, 7); - lean_ctor_release(x_7, 8); - lean_ctor_release(x_7, 9); - lean_ctor_release(x_7, 10); - x_139 = x_7; -} else { - lean_dec_ref(x_7); - x_139 = lean_box(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; +x_35 = lean_ctor_get(x_28, 0); +x_36 = lean_ctor_get(x_28, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_28); +x_37 = lean_ctor_get(x_35, 2); +lean_inc(x_37); +x_38 = lean_ctor_get(x_35, 4); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_40, 0, x_26); +lean_ctor_set(x_40, 1, x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_41, 1, x_36); +return x_41; } -x_140 = 1; -lean_inc(x_137); -lean_inc(x_136); -lean_inc(x_135); -lean_inc(x_134); -lean_inc(x_133); -lean_inc(x_132); -lean_inc(x_131); -lean_inc(x_130); -lean_inc(x_129); -lean_inc(x_128); -lean_inc(x_127); -if (lean_is_scalar(x_139)) { - x_141 = lean_alloc_ctor(0, 11, 2); -} else { - x_141 = x_139; -} -lean_ctor_set(x_141, 0, x_127); -lean_ctor_set(x_141, 1, x_128); -lean_ctor_set(x_141, 2, x_129); -lean_ctor_set(x_141, 3, x_130); -lean_ctor_set(x_141, 4, x_131); -lean_ctor_set(x_141, 5, x_132); -lean_ctor_set(x_141, 6, x_133); -lean_ctor_set(x_141, 7, x_134); -lean_ctor_set(x_141, 8, x_135); -lean_ctor_set(x_141, 9, x_136); -lean_ctor_set(x_141, 10, x_137); -lean_ctor_set_uint8(x_141, sizeof(void*)*11, x_140); -lean_ctor_set_uint8(x_141, sizeof(void*)*11 + 1, x_138); -x_142 = 0; -x_143 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_143, 0, x_127); -lean_ctor_set(x_143, 1, x_128); -lean_ctor_set(x_143, 2, x_129); -lean_ctor_set(x_143, 3, x_130); -lean_ctor_set(x_143, 4, x_131); -lean_ctor_set(x_143, 5, x_132); -lean_ctor_set(x_143, 6, x_133); -lean_ctor_set(x_143, 7, x_134); -lean_ctor_set(x_143, 8, x_135); -lean_ctor_set(x_143, 9, x_136); -lean_ctor_set(x_143, 10, x_137); -lean_ctor_set_uint8(x_143, sizeof(void*)*11, x_142); -lean_ctor_set_uint8(x_143, sizeof(void*)*11 + 1, x_138); -x_144 = lean_st_mk_ref(x_13, x_9); -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -lean_dec(x_144); -lean_inc(x_8); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_145); -x_147 = lean_dsimp(x_1, x_4, x_2, x_145, x_5, x_6, x_143, x_8, x_146); -if (lean_obj_tag(x_147) == 0) +} +else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -lean_dec(x_141); -lean_dec(x_5); -lean_dec(x_8); -lean_dec(x_6); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_150 = lean_st_ref_get(x_145, x_149); -lean_dec(x_145); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_153 = x_150; -} else { - lean_dec_ref(x_150); - x_153 = lean_box(0); +uint8_t x_42; +lean_dec(x_23); +x_42 = !lean_is_exclusive(x_25); +if (x_42 == 0) +{ +return x_25; +} +else +{ +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_25, 0); +x_44 = lean_ctor_get(x_25, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_25); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } -x_154 = lean_ctor_get(x_151, 2); -lean_inc(x_154); -lean_dec(x_151); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_148); -lean_ctor_set(x_155, 1, x_154); -if (lean_is_scalar(x_153)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_153; } -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_152); -return x_156; } else { -lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; -lean_dec(x_145); -x_157 = lean_ctor_get(x_147, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_147, 1); -lean_inc(x_158); -if (lean_is_exclusive(x_147)) { - lean_ctor_release(x_147, 0); - lean_ctor_release(x_147, 1); - x_159 = x_147; +uint8_t x_46; 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; lean_object* x_57; uint8_t x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_46 = lean_ctor_get_uint8(x_17, 0); +x_47 = lean_ctor_get_uint8(x_17, 1); +x_48 = lean_ctor_get_uint8(x_17, 2); +x_49 = lean_ctor_get_uint8(x_17, 3); +x_50 = lean_ctor_get_uint8(x_17, 4); +x_51 = lean_ctor_get_uint8(x_17, 5); +x_52 = lean_ctor_get_uint8(x_17, 6); +x_53 = lean_ctor_get_uint8(x_17, 7); +x_54 = lean_ctor_get_uint8(x_17, 8); +x_55 = lean_ctor_get_uint8(x_17, 10); +x_56 = lean_ctor_get_uint8(x_17, 12); +lean_dec(x_17); +x_57 = lean_ctor_get(x_2, 0); +lean_inc(x_57); +x_58 = lean_ctor_get_uint8(x_57, sizeof(void*)*2 + 6); +lean_dec(x_57); +x_59 = 2; +x_60 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_60, 0, x_46); +lean_ctor_set_uint8(x_60, 1, x_47); +lean_ctor_set_uint8(x_60, 2, x_48); +lean_ctor_set_uint8(x_60, 3, x_49); +lean_ctor_set_uint8(x_60, 4, x_50); +lean_ctor_set_uint8(x_60, 5, x_51); +lean_ctor_set_uint8(x_60, 6, x_52); +lean_ctor_set_uint8(x_60, 7, x_53); +lean_ctor_set_uint8(x_60, 8, x_54); +lean_ctor_set_uint8(x_60, 9, x_59); +lean_ctor_set_uint8(x_60, 10, x_55); +lean_ctor_set_uint8(x_60, 11, x_58); +lean_ctor_set_uint8(x_60, 12, x_56); +lean_ctor_set(x_5, 0, x_60); +x_61 = lean_st_mk_ref(x_15, x_9); +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_62); +x_64 = l_Lean_Meta_Simp_dsimpMain_dsimpMain(x_1, x_4, x_2, x_62, x_5, x_6, x_7, x_8, 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; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_65 = lean_ctor_get(x_64, 0); +lean_inc(x_65); +x_66 = lean_ctor_get(x_64, 1); +lean_inc(x_66); +lean_dec(x_64); +x_67 = lean_st_ref_get(x_62, x_66); +lean_dec(x_62); +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_67)) { + lean_ctor_release(x_67, 0); + lean_ctor_release(x_67, 1); + x_70 = x_67; } else { - lean_dec_ref(x_147); - x_159 = lean_box(0); + lean_dec_ref(x_67); + x_70 = lean_box(0); } -x_160 = l_Lean_Exception_isRuntime(x_157); -if (x_160 == 0) -{ -lean_object* x_161; -lean_dec(x_141); -lean_dec(x_5); -lean_dec(x_8); -lean_dec(x_6); -if (lean_is_scalar(x_159)) { - x_161 = lean_alloc_ctor(1, 2, 0); +x_71 = lean_ctor_get(x_68, 2); +lean_inc(x_71); +x_72 = lean_ctor_get(x_68, 4); +lean_inc(x_72); +lean_dec(x_68); +x_73 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_65); +lean_ctor_set(x_74, 1, x_73); +if (lean_is_scalar(x_70)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_161 = x_159; + x_75 = x_70; } -lean_ctor_set(x_161, 0, x_157); -lean_ctor_set(x_161, 1, x_158); -return x_161; +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_69); +return x_75; } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -lean_dec(x_159); -x_162 = l_Lean_Meta_Simp_dsimpMain___closed__2; -x_163 = l_Lean_Meta_throwNestedTacticEx___rarg(x_162, x_157, x_5, x_6, x_141, x_8, x_158); -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); -lean_inc(x_165); -if (lean_is_exclusive(x_163)) { - lean_ctor_release(x_163, 0); - lean_ctor_release(x_163, 1); - x_166 = x_163; +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_62); +x_76 = lean_ctor_get(x_64, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_64, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_78 = x_64; } else { - lean_dec_ref(x_163); - x_166 = lean_box(0); + lean_dec_ref(x_64); + x_78 = lean_box(0); } -if (lean_is_scalar(x_166)) { - x_167 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_167 = x_166; -} -lean_ctor_set(x_167, 0, x_164); -lean_ctor_set(x_167, 1, x_165); -return x_167; + x_79 = x_78; } +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; } } } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; uint8_t x_178; uint8_t x_179; uint8_t x_180; uint8_t x_181; uint8_t x_182; uint8_t x_183; uint8_t x_184; uint8_t x_185; uint8_t x_186; lean_object* x_187; lean_object* x_188; uint8_t x_189; uint8_t 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; uint8_t x_204; lean_object* x_205; uint8_t x_206; lean_object* x_207; uint8_t x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_168 = lean_ctor_get(x_5, 0); -x_169 = lean_ctor_get(x_5, 1); -x_170 = lean_ctor_get(x_5, 2); -x_171 = lean_ctor_get(x_5, 3); -x_172 = lean_ctor_get(x_5, 4); -x_173 = lean_ctor_get(x_5, 5); -x_174 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); -x_175 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); -lean_inc(x_173); -lean_inc(x_172); -lean_inc(x_171); -lean_inc(x_170); -lean_inc(x_169); -lean_inc(x_168); +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; 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; uint8_t x_94; uint8_t x_95; uint8_t x_96; uint8_t x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; uint8_t x_101; uint8_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_80 = lean_ctor_get(x_5, 0); +x_81 = lean_ctor_get(x_5, 1); +x_82 = lean_ctor_get(x_5, 2); +x_83 = lean_ctor_get(x_5, 3); +x_84 = lean_ctor_get(x_5, 4); +x_85 = lean_ctor_get(x_5, 5); +x_86 = lean_ctor_get_uint8(x_5, sizeof(void*)*6); +x_87 = lean_ctor_get_uint8(x_5, sizeof(void*)*6 + 1); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); lean_dec(x_5); -x_176 = lean_ctor_get_uint8(x_168, 0); -x_177 = lean_ctor_get_uint8(x_168, 1); -x_178 = lean_ctor_get_uint8(x_168, 2); -x_179 = lean_ctor_get_uint8(x_168, 3); -x_180 = lean_ctor_get_uint8(x_168, 4); -x_181 = lean_ctor_get_uint8(x_168, 5); -x_182 = lean_ctor_get_uint8(x_168, 6); -x_183 = lean_ctor_get_uint8(x_168, 7); -x_184 = lean_ctor_get_uint8(x_168, 8); -x_185 = lean_ctor_get_uint8(x_168, 10); -x_186 = lean_ctor_get_uint8(x_168, 12); -if (lean_is_exclusive(x_168)) { - x_187 = x_168; -} else { - lean_dec_ref(x_168); - x_187 = lean_box(0); -} -x_188 = lean_ctor_get(x_2, 0); -lean_inc(x_188); -x_189 = lean_ctor_get_uint8(x_188, sizeof(void*)*2 + 6); -lean_dec(x_188); -x_190 = 2; -if (lean_is_scalar(x_187)) { - x_191 = lean_alloc_ctor(0, 0, 13); -} else { - x_191 = x_187; -} -lean_ctor_set_uint8(x_191, 0, x_176); -lean_ctor_set_uint8(x_191, 1, x_177); -lean_ctor_set_uint8(x_191, 2, x_178); -lean_ctor_set_uint8(x_191, 3, x_179); -lean_ctor_set_uint8(x_191, 4, x_180); -lean_ctor_set_uint8(x_191, 5, x_181); -lean_ctor_set_uint8(x_191, 6, x_182); -lean_ctor_set_uint8(x_191, 7, x_183); -lean_ctor_set_uint8(x_191, 8, x_184); -lean_ctor_set_uint8(x_191, 9, x_190); -lean_ctor_set_uint8(x_191, 10, x_185); -lean_ctor_set_uint8(x_191, 11, x_189); -lean_ctor_set_uint8(x_191, 12, x_186); -x_192 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_169); -lean_ctor_set(x_192, 2, x_170); -lean_ctor_set(x_192, 3, x_171); -lean_ctor_set(x_192, 4, x_172); -lean_ctor_set(x_192, 5, x_173); -lean_ctor_set_uint8(x_192, sizeof(void*)*6, x_174); -lean_ctor_set_uint8(x_192, sizeof(void*)*6 + 1, x_175); -x_193 = lean_ctor_get(x_7, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_7, 1); -lean_inc(x_194); -x_195 = lean_ctor_get(x_7, 2); -lean_inc(x_195); -x_196 = lean_ctor_get(x_7, 3); -lean_inc(x_196); -x_197 = lean_ctor_get(x_7, 4); -lean_inc(x_197); -x_198 = lean_ctor_get(x_7, 5); -lean_inc(x_198); -x_199 = lean_ctor_get(x_7, 6); -lean_inc(x_199); -x_200 = lean_ctor_get(x_7, 7); -lean_inc(x_200); -x_201 = lean_ctor_get(x_7, 8); -lean_inc(x_201); -x_202 = lean_ctor_get(x_7, 9); -lean_inc(x_202); -x_203 = lean_ctor_get(x_7, 10); -lean_inc(x_203); -x_204 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -if (lean_is_exclusive(x_7)) { - lean_ctor_release(x_7, 0); - lean_ctor_release(x_7, 1); - lean_ctor_release(x_7, 2); - lean_ctor_release(x_7, 3); - lean_ctor_release(x_7, 4); - lean_ctor_release(x_7, 5); - lean_ctor_release(x_7, 6); - lean_ctor_release(x_7, 7); - lean_ctor_release(x_7, 8); - lean_ctor_release(x_7, 9); - lean_ctor_release(x_7, 10); - x_205 = x_7; +x_88 = lean_ctor_get_uint8(x_80, 0); +x_89 = lean_ctor_get_uint8(x_80, 1); +x_90 = lean_ctor_get_uint8(x_80, 2); +x_91 = lean_ctor_get_uint8(x_80, 3); +x_92 = lean_ctor_get_uint8(x_80, 4); +x_93 = lean_ctor_get_uint8(x_80, 5); +x_94 = lean_ctor_get_uint8(x_80, 6); +x_95 = lean_ctor_get_uint8(x_80, 7); +x_96 = lean_ctor_get_uint8(x_80, 8); +x_97 = lean_ctor_get_uint8(x_80, 10); +x_98 = lean_ctor_get_uint8(x_80, 12); +if (lean_is_exclusive(x_80)) { + x_99 = x_80; } else { - lean_dec_ref(x_7); - x_205 = lean_box(0); + lean_dec_ref(x_80); + x_99 = lean_box(0); } -x_206 = 1; -lean_inc(x_203); -lean_inc(x_202); -lean_inc(x_201); -lean_inc(x_200); -lean_inc(x_199); -lean_inc(x_198); -lean_inc(x_197); -lean_inc(x_196); -lean_inc(x_195); -lean_inc(x_194); -lean_inc(x_193); -if (lean_is_scalar(x_205)) { - x_207 = lean_alloc_ctor(0, 11, 2); -} else { - x_207 = x_205; -} -lean_ctor_set(x_207, 0, x_193); -lean_ctor_set(x_207, 1, x_194); -lean_ctor_set(x_207, 2, x_195); -lean_ctor_set(x_207, 3, x_196); -lean_ctor_set(x_207, 4, x_197); -lean_ctor_set(x_207, 5, x_198); -lean_ctor_set(x_207, 6, x_199); -lean_ctor_set(x_207, 7, x_200); -lean_ctor_set(x_207, 8, x_201); -lean_ctor_set(x_207, 9, x_202); -lean_ctor_set(x_207, 10, x_203); -lean_ctor_set_uint8(x_207, sizeof(void*)*11, x_206); -lean_ctor_set_uint8(x_207, sizeof(void*)*11 + 1, x_204); -x_208 = 0; -x_209 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_209, 0, x_193); -lean_ctor_set(x_209, 1, x_194); -lean_ctor_set(x_209, 2, x_195); -lean_ctor_set(x_209, 3, x_196); -lean_ctor_set(x_209, 4, x_197); -lean_ctor_set(x_209, 5, x_198); -lean_ctor_set(x_209, 6, x_199); -lean_ctor_set(x_209, 7, x_200); -lean_ctor_set(x_209, 8, x_201); -lean_ctor_set(x_209, 9, x_202); -lean_ctor_set(x_209, 10, x_203); -lean_ctor_set_uint8(x_209, sizeof(void*)*11, x_208); -lean_ctor_set_uint8(x_209, sizeof(void*)*11 + 1, x_204); -x_210 = lean_st_mk_ref(x_13, x_9); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -lean_inc(x_8); -lean_inc(x_6); -lean_inc(x_192); -lean_inc(x_211); -x_213 = lean_dsimp(x_1, x_4, x_2, x_211, x_192, x_6, x_209, x_8, x_212); -if (lean_obj_tag(x_213) == 0) +x_100 = lean_ctor_get(x_2, 0); +lean_inc(x_100); +x_101 = lean_ctor_get_uint8(x_100, sizeof(void*)*2 + 6); +lean_dec(x_100); +x_102 = 2; +if (lean_is_scalar(x_99)) { + x_103 = lean_alloc_ctor(0, 0, 13); +} else { + x_103 = x_99; +} +lean_ctor_set_uint8(x_103, 0, x_88); +lean_ctor_set_uint8(x_103, 1, x_89); +lean_ctor_set_uint8(x_103, 2, x_90); +lean_ctor_set_uint8(x_103, 3, x_91); +lean_ctor_set_uint8(x_103, 4, x_92); +lean_ctor_set_uint8(x_103, 5, x_93); +lean_ctor_set_uint8(x_103, 6, x_94); +lean_ctor_set_uint8(x_103, 7, x_95); +lean_ctor_set_uint8(x_103, 8, x_96); +lean_ctor_set_uint8(x_103, 9, x_102); +lean_ctor_set_uint8(x_103, 10, x_97); +lean_ctor_set_uint8(x_103, 11, x_101); +lean_ctor_set_uint8(x_103, 12, x_98); +x_104 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_81); +lean_ctor_set(x_104, 2, x_82); +lean_ctor_set(x_104, 3, x_83); +lean_ctor_set(x_104, 4, x_84); +lean_ctor_set(x_104, 5, x_85); +lean_ctor_set_uint8(x_104, sizeof(void*)*6, x_86); +lean_ctor_set_uint8(x_104, sizeof(void*)*6 + 1, x_87); +x_105 = lean_st_mk_ref(x_15, x_9); +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); +lean_inc(x_106); +x_108 = l_Lean_Meta_Simp_dsimpMain_dsimpMain(x_1, x_4, x_2, x_106, x_104, x_6, x_7, x_8, x_107); +if (lean_obj_tag(x_108) == 0) { -lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -lean_dec(x_207); -lean_dec(x_192); -lean_dec(x_8); -lean_dec(x_6); -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); -lean_inc(x_215); -lean_dec(x_213); -x_216 = lean_st_ref_get(x_211, x_215); -lean_dec(x_211); -x_217 = lean_ctor_get(x_216, 0); -lean_inc(x_217); -x_218 = lean_ctor_get(x_216, 1); -lean_inc(x_218); -if (lean_is_exclusive(x_216)) { - lean_ctor_release(x_216, 0); - lean_ctor_release(x_216, 1); - x_219 = x_216; +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_st_ref_get(x_106, x_110); +lean_dec(x_106); +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; } else { - lean_dec_ref(x_216); - x_219 = lean_box(0); + lean_dec_ref(x_111); + x_114 = lean_box(0); } -x_220 = lean_ctor_get(x_217, 2); -lean_inc(x_220); -lean_dec(x_217); -x_221 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_221, 0, x_214); -lean_ctor_set(x_221, 1, x_220); -if (lean_is_scalar(x_219)) { - x_222 = lean_alloc_ctor(0, 2, 0); +x_115 = lean_ctor_get(x_112, 2); +lean_inc(x_115); +x_116 = lean_ctor_get(x_112, 4); +lean_inc(x_116); +lean_dec(x_112); +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_115); +lean_ctor_set(x_117, 1, x_116); +x_118 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_117); +if (lean_is_scalar(x_114)) { + x_119 = lean_alloc_ctor(0, 2, 0); } else { - x_222 = x_219; + x_119 = x_114; } -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_218); -return x_222; +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_113); +return x_119; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; -lean_dec(x_211); -x_223 = lean_ctor_get(x_213, 0); -lean_inc(x_223); -x_224 = lean_ctor_get(x_213, 1); -lean_inc(x_224); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_225 = x_213; -} else { - lean_dec_ref(x_213); - x_225 = lean_box(0); -} -x_226 = l_Lean_Exception_isRuntime(x_223); -if (x_226 == 0) -{ -lean_object* x_227; -lean_dec(x_207); -lean_dec(x_192); -lean_dec(x_8); -lean_dec(x_6); -if (lean_is_scalar(x_225)) { - x_227 = lean_alloc_ctor(1, 2, 0); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_106); +x_120 = lean_ctor_get(x_108, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_108, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_122 = x_108; } else { - x_227 = x_225; -} -lean_ctor_set(x_227, 0, x_223); -lean_ctor_set(x_227, 1, x_224); -return x_227; + lean_dec_ref(x_108); + x_122 = lean_box(0); } -else -{ -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_dec(x_225); -x_228 = l_Lean_Meta_Simp_dsimpMain___closed__2; -x_229 = l_Lean_Meta_throwNestedTacticEx___rarg(x_228, x_223, x_192, x_6, x_207, x_8, x_224); -x_230 = lean_ctor_get(x_229, 0); -lean_inc(x_230); -x_231 = lean_ctor_get(x_229, 1); -lean_inc(x_231); -if (lean_is_exclusive(x_229)) { - lean_ctor_release(x_229, 0); - lean_ctor_release(x_229, 1); - x_232 = x_229; +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); } else { - lean_dec_ref(x_229); - x_232 = lean_box(0); + x_123 = x_122; } -if (lean_is_scalar(x_232)) { - x_233 = lean_alloc_ctor(1, 2, 0); -} else { - x_233 = x_232; +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; } -lean_ctor_set(x_233, 0, x_230); -lean_ctor_set(x_233, 1, x_231); -return x_233; } } } +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimpMain___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Meta_Simp_dsimpMain(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_3); +return x_10; } } static lean_object* _init_l_Lean_Meta_simp___closed__1() { @@ -34003,62 +34021,73 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simp(lean_object* x_1, lean_object* x_2, le { if (lean_obj_tag(x_4) == 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_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_11 = lean_ctor_get(x_8, 2); lean_inc(x_11); x_12 = l_Lean_Meta_simp___closed__1; -x_13 = l_Lean_Meta_Simp_mkMethods(x_3, x_12); -x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_main), 9, 4); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_2); -lean_closure_set(x_14, 2, x_5); -lean_closure_set(x_14, 3, x_13); -x_15 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; -x_16 = lean_box(0); -x_17 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_15, x_11, x_14, x_16, x_6, x_7, x_8, x_9, x_10); +x_13 = 1; +x_14 = l_Lean_Meta_Simp_mkMethods(x_3, x_12, x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_main___boxed), 9, 4); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +lean_closure_set(x_15, 2, x_5); +lean_closure_set(x_15, 3, x_14); +x_16 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; +x_17 = lean_box(0); +x_18 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__15___rarg(x_16, x_11, x_15, x_17, x_6, x_7, x_8, x_9, x_10); lean_dec(x_11); -return x_17; +return x_18; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_18 = lean_ctor_get(x_8, 2); -lean_inc(x_18); -x_19 = lean_ctor_get(x_4, 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; lean_object* x_26; +x_19 = lean_ctor_get(x_8, 2); lean_inc(x_19); +x_20 = lean_ctor_get(x_4, 0); +lean_inc(x_20); lean_dec(x_4); -x_20 = l_Lean_Meta_Simp_mkMethods(x_3, x_19); -x_21 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_main), 9, 4); -lean_closure_set(x_21, 0, x_1); -lean_closure_set(x_21, 1, x_2); -lean_closure_set(x_21, 2, x_5); -lean_closure_set(x_21, 3, x_20); -x_22 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; -x_23 = lean_box(0); -x_24 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_22, x_18, x_21, x_23, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_18); -return x_24; +x_21 = 0; +x_22 = l_Lean_Meta_Simp_mkMethods(x_3, x_20, x_21); +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_main___boxed), 9, 4); +lean_closure_set(x_23, 0, x_1); +lean_closure_set(x_23, 1, x_2); +lean_closure_set(x_23, 2, x_5); +lean_closure_set(x_23, 3, x_22); +x_24 = l___private_Lean_Meta_Tactic_Simp_Main_0__Lean_Meta_Simp_reduceStep___lambda__2___closed__3; +x_25 = lean_box(0); +x_26 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__15___rarg(x_24, x_19, x_23, x_25, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_19); +return x_26; +} +} } +static lean_object* _init_l_Lean_Meta_dsimp___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("dsimp", 5); +return x_1; } } LEAN_EXPORT lean_object* l_Lean_Meta_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) { _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_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; x_10 = lean_ctor_get(x_7, 2); lean_inc(x_10); x_11 = l_Lean_Meta_simp___closed__1; -x_12 = l_Lean_Meta_Simp_mkMethods(x_3, x_11); -x_13 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dsimpMain), 9, 4); -lean_closure_set(x_13, 0, x_1); -lean_closure_set(x_13, 1, x_2); -lean_closure_set(x_13, 2, x_4); -lean_closure_set(x_13, 3, x_12); -x_14 = l_Lean_Meta_Simp_dsimpMain___closed__1; -x_15 = lean_box(0); -x_16 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__16___rarg(x_14, x_10, x_13, x_15, x_5, x_6, x_7, x_8, x_9); +x_12 = 1; +x_13 = l_Lean_Meta_Simp_mkMethods(x_3, x_11, x_12); +x_14 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_dsimpMain___boxed), 9, 4); +lean_closure_set(x_14, 0, x_1); +lean_closure_set(x_14, 1, x_2); +lean_closure_set(x_14, 2, x_4); +lean_closure_set(x_14, 3, x_13); +x_15 = l_Lean_Meta_dsimp___closed__1; +x_16 = lean_box(0); +x_17 = l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___spec__15___rarg(x_15, x_10, x_14, x_16, x_5, x_6, x_7, x_8, x_9); lean_dec(x_10); -return x_16; +return x_17; } } LEAN_EXPORT lean_object* l_Lean_Meta_simpTargetCore___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) { @@ -34747,7 +34776,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simpTarget(lean_object* x_1, lean_object* x _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l_Lean_Meta_Simp_main___closed__1; +x_12 = l_Lean_Meta_Simp_main_simpMain___closed__1; x_13 = lean_box(x_5); lean_inc(x_1); x_14 = lean_alloc_closure((void*)(l_Lean_Meta_simpTarget___lambda__1___boxed), 12, 7); @@ -37277,7 +37306,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simpLocalDecl(lean_object* x_1, lean_object _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = l_Lean_Meta_Simp_main___closed__1; +x_13 = l_Lean_Meta_Simp_main_simpMain___closed__1; x_14 = lean_box(x_6); lean_inc(x_1); x_15 = lean_alloc_closure((void*)(l_Lean_Meta_simpLocalDecl___lambda__1___boxed), 13, 8); @@ -37463,7 +37492,7 @@ lean_inc(x_16); x_38 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_38) == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint32_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint32_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint32_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; uint32_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; x_39 = lean_ctor_get(x_38, 0); lean_inc(x_39); x_40 = lean_ctor_get(x_38, 1); @@ -37478,25 +37507,28 @@ lean_inc(x_44); lean_dec(x_42); x_45 = lean_ctor_get(x_1, 0); lean_inc(x_45); -x_46 = lean_ctor_get_uint32(x_1, sizeof(void*)*4); +x_46 = lean_ctor_get_uint32(x_1, sizeof(void*)*5); x_47 = lean_ctor_get(x_1, 1); lean_inc(x_47); x_48 = lean_ctor_get(x_1, 2); lean_inc(x_48); x_49 = lean_ctor_get(x_1, 3); lean_inc(x_49); -x_50 = lean_ctor_get_uint32(x_1, sizeof(void*)*4 + 4); -x_51 = l_Lean_LocalDecl_fvarId(x_39); -x_52 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_52, 0, x_51); -x_53 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_47, x_52); -x_54 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_54, 0, x_45); -lean_ctor_set(x_54, 1, x_53); -lean_ctor_set(x_54, 2, x_48); -lean_ctor_set(x_54, 3, x_49); -lean_ctor_set_uint32(x_54, sizeof(void*)*4, x_46); -lean_ctor_set_uint32(x_54, sizeof(void*)*4 + 4, x_50); +x_50 = lean_ctor_get_uint32(x_1, sizeof(void*)*5 + 4); +x_51 = lean_ctor_get(x_1, 4); +lean_inc(x_51); +x_52 = l_Lean_LocalDecl_fvarId(x_39); +x_53 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_53, 0, x_52); +x_54 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_47, x_53); +x_55 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_55, 0, x_45); +lean_ctor_set(x_55, 1, x_54); +lean_ctor_set(x_55, 2, x_48); +lean_ctor_set(x_55, 3, x_49); +lean_ctor_set(x_55, 4, x_51); +lean_ctor_set_uint32(x_55, sizeof(void*)*5, x_46); +lean_ctor_set_uint32(x_55, sizeof(void*)*5 + 4, x_50); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); @@ -37504,63 +37536,63 @@ lean_inc(x_9); lean_inc(x_3); lean_inc(x_2); lean_inc(x_43); -x_55 = l_Lean_Meta_simp(x_43, x_54, x_2, x_3, x_37, x_9, x_10, x_11, x_12, x_44); -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_55, 0); -lean_inc(x_56); +x_56 = l_Lean_Meta_simp(x_43, x_55, x_2, x_3, x_36, x_9, x_10, x_11, x_12, x_44); +if (lean_obj_tag(x_56) == 0) +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; x_57 = lean_ctor_get(x_56, 0); lean_inc(x_57); -x_58 = lean_ctor_get(x_57, 1); +x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); -if (lean_obj_tag(x_58) == 0) +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_59; uint8_t x_60; +lean_object* x_60; uint8_t x_61; lean_dec(x_43); lean_dec(x_39); -x_59 = lean_ctor_get(x_55, 1); -lean_inc(x_59); -lean_dec(x_55); -x_60 = !lean_is_exclusive(x_56); -if (x_60 == 0) +x_60 = lean_ctor_get(x_56, 1); +lean_inc(x_60); +lean_dec(x_56); +x_61 = !lean_is_exclusive(x_57); +if (x_61 == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -x_61 = lean_ctor_get(x_56, 1); -x_62 = lean_ctor_get(x_56, 0); -lean_dec(x_62); +lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; +x_62 = lean_ctor_get(x_57, 1); x_63 = lean_ctor_get(x_57, 0); -lean_inc(x_63); -x_64 = l_Lean_Expr_isFalse(x_63); -if (x_64 == 0) +lean_dec(x_63); +x_64 = lean_ctor_get(x_58, 0); +lean_inc(x_64); +x_65 = l_Lean_Expr_isFalse(x_64); +if (x_65 == 0) { -lean_object* x_65; lean_object* x_66; -lean_free_object(x_56); +lean_object* x_66; lean_object* x_67; +lean_free_object(x_57); lean_free_object(x_28); lean_free_object(x_27); lean_free_object(x_26); -x_65 = lean_box(0); +x_66 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_4); -x_66 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_57, x_16, x_36, x_61, x_4, x_30, x_33, x_65, x_9, x_10, x_11, x_12, x_59); -if (lean_obj_tag(x_66) == 0) +x_67 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_58, x_16, x_62, x_37, x_4, x_30, x_33, x_66, x_9, x_10, x_11, x_12, x_60); +if (lean_obj_tag(x_67) == 0) { -lean_object* x_67; lean_object* x_68; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +lean_object* x_68; lean_object* x_69; +x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); -lean_dec(x_66); -x_17 = x_67; -x_18 = x_68; +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); +lean_dec(x_67); +x_17 = x_68; +x_18 = x_69; goto block_25; } else { -uint8_t x_69; +uint8_t x_70; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -37569,81 +37601,81 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_69 = !lean_is_exclusive(x_66); -if (x_69 == 0) +x_70 = !lean_is_exclusive(x_67); +if (x_70 == 0) { -return x_66; +return x_67; } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; -x_70 = lean_ctor_get(x_66, 0); -x_71 = lean_ctor_get(x_66, 1); +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_67, 0); +x_72 = lean_ctor_get(x_67, 1); +lean_inc(x_72); lean_inc(x_71); -lean_inc(x_70); -lean_dec(x_66); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_71); -return x_72; +lean_dec(x_67); +x_73 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +return x_73; } } } else { -lean_object* x_73; -lean_dec(x_57); +lean_object* x_74; +lean_dec(x_58); lean_inc(x_30); -x_73 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_59); -if (lean_obj_tag(x_73) == 0) +x_74 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_60); +if (lean_obj_tag(x_74) == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_74 = lean_ctor_get(x_73, 0); -lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); +lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_75 = lean_ctor_get(x_74, 0); lean_inc(x_75); -lean_dec(x_73); -x_76 = l_Lean_Expr_fvar___override(x_16); +x_76 = lean_ctor_get(x_74, 1); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Expr_fvar___override(x_16); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_77 = l_Lean_Meta_mkFalseElim(x_74, x_76, x_9, x_10, x_11, x_12, x_75); -if (lean_obj_tag(x_77) == 0) +x_78 = l_Lean_Meta_mkFalseElim(x_75, x_77, x_9, x_10, x_11, x_12, x_76); +if (lean_obj_tag(x_78) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_79 = lean_ctor_get(x_78, 0); lean_inc(x_79); -lean_dec(x_77); +x_80 = lean_ctor_get(x_78, 1); +lean_inc(x_80); +lean_dec(x_78); lean_inc(x_30); -x_80 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_78, x_9, x_10, x_11, x_12, x_79); -x_81 = lean_ctor_get(x_80, 1); -lean_inc(x_81); -lean_dec(x_80); -x_82 = lean_box(0); -lean_inc(x_61); -lean_ctor_set(x_56, 0, x_82); -x_83 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_83, 0, x_56); -lean_ctor_set(x_28, 1, x_61); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_26); -x_85 = lean_alloc_ctor(0, 1, 0); +x_81 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_79, x_9, x_10, x_11, x_12, x_80); +x_82 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_81); +x_83 = lean_box(0); +lean_inc(x_62); +lean_ctor_set(x_57, 0, x_83); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_57); +lean_ctor_set(x_28, 0, x_62); +x_85 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_85, 0, x_84); -x_17 = x_85; -x_18 = x_81; +lean_ctor_set(x_85, 1, x_26); +x_86 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_86, 0, x_85); +x_17 = x_86; +x_18 = x_82; goto block_25; } else { -uint8_t x_86; -lean_free_object(x_56); -lean_dec(x_61); +uint8_t x_87; +lean_free_object(x_57); +lean_dec(x_62); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -37656,33 +37688,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_86 = !lean_is_exclusive(x_77); -if (x_86 == 0) +x_87 = !lean_is_exclusive(x_78); +if (x_87 == 0) { -return x_77; +return x_78; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_77, 0); -x_88 = lean_ctor_get(x_77, 1); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_78, 0); +x_89 = lean_ctor_get(x_78, 1); +lean_inc(x_89); lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_77); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_dec(x_78); +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 { -uint8_t x_90; -lean_free_object(x_56); -lean_dec(x_61); +uint8_t x_91; +lean_free_object(x_57); +lean_dec(x_62); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -37696,64 +37728,64 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_90 = !lean_is_exclusive(x_73); -if (x_90 == 0) +x_91 = !lean_is_exclusive(x_74); +if (x_91 == 0) { -return x_73; +return x_74; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_73, 0); -x_92 = lean_ctor_get(x_73, 1); +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_74, 0); +x_93 = lean_ctor_get(x_74, 1); +lean_inc(x_93); lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_73); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_dec(x_74); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; } } } } else { -lean_object* x_94; lean_object* x_95; uint8_t x_96; -x_94 = lean_ctor_get(x_56, 1); -lean_inc(x_94); -lean_dec(x_56); -x_95 = lean_ctor_get(x_57, 0); +lean_object* x_95; lean_object* x_96; uint8_t x_97; +x_95 = lean_ctor_get(x_57, 1); lean_inc(x_95); -x_96 = l_Lean_Expr_isFalse(x_95); -if (x_96 == 0) +lean_dec(x_57); +x_96 = lean_ctor_get(x_58, 0); +lean_inc(x_96); +x_97 = l_Lean_Expr_isFalse(x_96); +if (x_97 == 0) { -lean_object* x_97; lean_object* x_98; +lean_object* x_98; lean_object* x_99; lean_free_object(x_28); lean_free_object(x_27); lean_free_object(x_26); -x_97 = lean_box(0); +x_98 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_4); -x_98 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_57, x_16, x_36, x_94, x_4, x_30, x_33, x_97, x_9, x_10, x_11, x_12, x_59); -if (lean_obj_tag(x_98) == 0) +x_99 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_58, x_16, x_95, x_37, x_4, x_30, x_33, x_98, x_9, x_10, x_11, x_12, x_60); +if (lean_obj_tag(x_99) == 0) { -lean_object* x_99; lean_object* x_100; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); +lean_object* x_100; lean_object* x_101; +x_100 = lean_ctor_get(x_99, 0); lean_inc(x_100); -lean_dec(x_98); -x_17 = x_99; -x_18 = x_100; +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_17 = x_100; +x_18 = x_101; goto block_25; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -37762,84 +37794,84 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_101 = lean_ctor_get(x_98, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_98, 1); +x_102 = lean_ctor_get(x_99, 0); lean_inc(x_102); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_103 = x_98; +x_103 = lean_ctor_get(x_99, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_104 = x_99; } else { - lean_dec_ref(x_98); - x_103 = lean_box(0); + lean_dec_ref(x_99); + x_104 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_105 = x_104; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } else { -lean_object* x_105; -lean_dec(x_57); +lean_object* x_106; +lean_dec(x_58); lean_inc(x_30); -x_105 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_59); -if (lean_obj_tag(x_105) == 0) +x_106 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_60); +if (lean_obj_tag(x_106) == 0) { -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_105, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_107 = lean_ctor_get(x_106, 0); lean_inc(x_107); -lean_dec(x_105); -x_108 = l_Lean_Expr_fvar___override(x_16); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = l_Lean_Expr_fvar___override(x_16); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_109 = l_Lean_Meta_mkFalseElim(x_106, x_108, x_9, x_10, x_11, x_12, x_107); -if (lean_obj_tag(x_109) == 0) +x_110 = l_Lean_Meta_mkFalseElim(x_107, x_109, x_9, x_10, x_11, x_12, x_108); +if (lean_obj_tag(x_110) == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_111 = lean_ctor_get(x_110, 0); lean_inc(x_111); -lean_dec(x_109); +x_112 = lean_ctor_get(x_110, 1); +lean_inc(x_112); +lean_dec(x_110); lean_inc(x_30); -x_112 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_110, x_9, x_10, x_11, x_12, x_111); -x_113 = lean_ctor_get(x_112, 1); -lean_inc(x_113); -lean_dec(x_112); -x_114 = lean_box(0); -lean_inc(x_94); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_94); -x_116 = lean_alloc_ctor(1, 1, 0); +x_113 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_111, x_9, x_10, x_11, x_12, x_112); +x_114 = lean_ctor_get(x_113, 1); +lean_inc(x_114); +lean_dec(x_113); +x_115 = lean_box(0); +lean_inc(x_95); +x_116 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_28, 1, x_94); -x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 1, x_95); +x_117 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_26); -x_118 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_28, 0, x_95); +x_118 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_118, 0, x_117); -x_17 = x_118; -x_18 = x_113; +lean_ctor_set(x_118, 1, x_26); +x_119 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_119, 0, x_118); +x_17 = x_119; +x_18 = x_114; goto block_25; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -lean_dec(x_94); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_95); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -37852,34 +37884,34 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_119 = lean_ctor_get(x_109, 0); -lean_inc(x_119); -x_120 = lean_ctor_get(x_109, 1); +x_120 = lean_ctor_get(x_110, 0); lean_inc(x_120); -if (lean_is_exclusive(x_109)) { - lean_ctor_release(x_109, 0); - lean_ctor_release(x_109, 1); - x_121 = x_109; +x_121 = lean_ctor_get(x_110, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + x_122 = x_110; } else { - lean_dec_ref(x_109); - x_121 = lean_box(0); + lean_dec_ref(x_110); + x_122 = lean_box(0); } -if (lean_is_scalar(x_121)) { - x_122 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); } else { - x_122 = x_121; + x_123 = x_122; } -lean_ctor_set(x_122, 0, x_119); -lean_ctor_set(x_122, 1, x_120); -return x_122; +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_94); +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_95); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -37893,129 +37925,129 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_123 = lean_ctor_get(x_105, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_105, 1); +x_124 = lean_ctor_get(x_106, 0); lean_inc(x_124); -if (lean_is_exclusive(x_105)) { - lean_ctor_release(x_105, 0); - lean_ctor_release(x_105, 1); - x_125 = x_105; +x_125 = lean_ctor_get(x_106, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_106)) { + lean_ctor_release(x_106, 0); + lean_ctor_release(x_106, 1); + x_126 = x_106; } else { - lean_dec_ref(x_105); - x_125 = lean_box(0); + lean_dec_ref(x_106); + x_126 = lean_box(0); } -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 2, 0); } else { - x_126 = x_125; + x_127 = x_126; } -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; +lean_ctor_set(x_127, 0, x_124); +lean_ctor_set(x_127, 1, x_125); +return x_127; } } } } else { -uint8_t x_127; -x_127 = !lean_is_exclusive(x_58); -if (x_127 == 0) +uint8_t x_128; +x_128 = !lean_is_exclusive(x_59); +if (x_128 == 0) { -lean_object* x_128; lean_object* x_129; uint8_t x_130; -x_128 = lean_ctor_get(x_58, 0); -lean_dec(x_128); -x_129 = lean_ctor_get(x_55, 1); -lean_inc(x_129); -lean_dec(x_55); -x_130 = !lean_is_exclusive(x_56); -if (x_130 == 0) +lean_object* x_129; lean_object* x_130; uint8_t x_131; +x_129 = lean_ctor_get(x_59, 0); +lean_dec(x_129); +x_130 = lean_ctor_get(x_56, 1); +lean_inc(x_130); +lean_dec(x_56); +x_131 = !lean_is_exclusive(x_57); +if (x_131 == 0) { -lean_object* x_131; lean_object* x_132; lean_object* x_133; uint8_t x_134; lean_object* x_135; -x_131 = lean_ctor_get(x_56, 1); -x_132 = lean_ctor_get(x_56, 0); -lean_dec(x_132); -x_133 = l_Lean_Expr_fvar___override(x_16); -x_134 = 1; +lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; +x_132 = lean_ctor_get(x_57, 1); +x_133 = lean_ctor_get(x_57, 0); +lean_dec(x_133); +x_134 = l_Lean_Expr_fvar___override(x_16); +x_135 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_30); -x_135 = l_Lean_Meta_applySimpResultToProp(x_30, x_133, x_43, x_57, x_134, x_9, x_10, x_11, x_12, x_129); +x_136 = l_Lean_Meta_applySimpResultToProp(x_30, x_134, x_43, x_58, x_135, x_9, x_10, x_11, x_12, x_130); lean_dec(x_43); -if (lean_obj_tag(x_135) == 0) -{ -lean_object* x_136; -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); if (lean_obj_tag(x_136) == 0) { -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; -lean_dec(x_39); -x_137 = lean_ctor_get(x_135, 1); +lean_object* x_137; +x_137 = lean_ctor_get(x_136, 0); lean_inc(x_137); -lean_dec(x_135); -x_138 = lean_box(0); -lean_inc(x_131); -lean_ctor_set(x_56, 0, x_138); -lean_ctor_set(x_58, 0, x_56); -lean_ctor_set(x_28, 1, x_131); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_58); -lean_ctor_set(x_139, 1, x_26); -x_140 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_140, 0, x_139); -x_17 = x_140; -x_18 = x_137; +if (lean_obj_tag(x_137) == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; +lean_dec(x_39); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +lean_dec(x_136); +x_139 = lean_box(0); +lean_inc(x_132); +lean_ctor_set(x_57, 0, x_139); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_28, 0, x_132); +x_140 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_140, 0, x_59); +lean_ctor_set(x_140, 1, x_26); +x_141 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_141, 0, x_140); +x_17 = x_141; +x_18 = x_138; goto block_25; } else { -lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -lean_free_object(x_56); -lean_free_object(x_58); -x_141 = lean_ctor_get(x_136, 0); -lean_inc(x_141); -lean_dec(x_136); -x_142 = lean_ctor_get(x_135, 1); +lean_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_free_object(x_57); +lean_free_object(x_59); +x_142 = lean_ctor_get(x_137, 0); lean_inc(x_142); -lean_dec(x_135); -x_143 = lean_ctor_get(x_141, 0); +lean_dec(x_137); +x_143 = lean_ctor_get(x_136, 1); lean_inc(x_143); -x_144 = lean_ctor_get(x_141, 1); +lean_dec(x_136); +x_144 = lean_ctor_get(x_142, 0); lean_inc(x_144); -lean_dec(x_141); -x_145 = l_Lean_LocalDecl_userName(x_39); +x_145 = lean_ctor_get(x_142, 1); +lean_inc(x_145); +lean_dec(x_142); +x_146 = l_Lean_LocalDecl_userName(x_39); lean_dec(x_39); -x_146 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_144); -lean_ctor_set(x_146, 2, x_143); -x_147 = lean_array_push(x_36, x_146); -lean_ctor_set(x_28, 1, x_131); -lean_ctor_set(x_28, 0, x_147); +x_147 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_147, 0, x_146); +lean_ctor_set(x_147, 1, x_145); +lean_ctor_set(x_147, 2, x_144); +x_148 = lean_array_push(x_37, x_147); +lean_ctor_set(x_28, 1, x_148); +lean_ctor_set(x_28, 0, x_132); lean_inc(x_4); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_4); -lean_ctor_set(x_148, 1, x_26); -x_149 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_149, 0, x_148); -x_17 = x_149; -x_18 = x_142; +x_149 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_149, 0, x_4); +lean_ctor_set(x_149, 1, x_26); +x_150 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_150, 0, x_149); +x_17 = x_150; +x_18 = x_143; goto block_25; } } else { -uint8_t x_150; -lean_free_object(x_56); -lean_dec(x_131); -lean_free_object(x_58); +uint8_t x_151; +lean_free_object(x_57); +lean_dec(x_132); +lean_free_object(x_59); lean_dec(x_39); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38028,112 +38060,112 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_150 = !lean_is_exclusive(x_135); -if (x_150 == 0) +x_151 = !lean_is_exclusive(x_136); +if (x_151 == 0) { -return x_135; +return x_136; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_135, 0); -x_152 = lean_ctor_get(x_135, 1); +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_136, 0); +x_153 = lean_ctor_get(x_136, 1); +lean_inc(x_153); lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_135); -x_153 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -return x_153; +lean_dec(x_136); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; } } } else { -lean_object* x_154; lean_object* x_155; uint8_t x_156; lean_object* x_157; -x_154 = lean_ctor_get(x_56, 1); -lean_inc(x_154); -lean_dec(x_56); -x_155 = l_Lean_Expr_fvar___override(x_16); -x_156 = 1; +lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; +x_155 = lean_ctor_get(x_57, 1); +lean_inc(x_155); +lean_dec(x_57); +x_156 = l_Lean_Expr_fvar___override(x_16); +x_157 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_30); -x_157 = l_Lean_Meta_applySimpResultToProp(x_30, x_155, x_43, x_57, x_156, x_9, x_10, x_11, x_12, x_129); +x_158 = l_Lean_Meta_applySimpResultToProp(x_30, x_156, x_43, x_58, x_157, x_9, x_10, x_11, x_12, x_130); lean_dec(x_43); -if (lean_obj_tag(x_157) == 0) -{ -lean_object* x_158; -x_158 = lean_ctor_get(x_157, 0); -lean_inc(x_158); if (lean_obj_tag(x_158) == 0) { -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_39); -x_159 = lean_ctor_get(x_157, 1); +lean_object* x_159; +x_159 = lean_ctor_get(x_158, 0); lean_inc(x_159); -lean_dec(x_157); -x_160 = lean_box(0); -lean_inc(x_154); -x_161 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_154); -lean_ctor_set(x_58, 0, x_161); -lean_ctor_set(x_28, 1, x_154); +if (lean_obj_tag(x_159) == 0) +{ +lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; +lean_dec(x_39); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +lean_dec(x_158); +x_161 = lean_box(0); +lean_inc(x_155); x_162 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_162, 0, x_58); -lean_ctor_set(x_162, 1, x_26); -x_163 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_163, 0, x_162); -x_17 = x_163; -x_18 = x_159; +lean_ctor_set(x_162, 0, x_161); +lean_ctor_set(x_162, 1, x_155); +lean_ctor_set(x_59, 0, x_162); +lean_ctor_set(x_28, 0, x_155); +x_163 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_163, 0, x_59); +lean_ctor_set(x_163, 1, x_26); +x_164 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_164, 0, x_163); +x_17 = x_164; +x_18 = x_160; goto block_25; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; -lean_free_object(x_58); -x_164 = lean_ctor_get(x_158, 0); -lean_inc(x_164); -lean_dec(x_158); -x_165 = lean_ctor_get(x_157, 1); +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; +lean_free_object(x_59); +x_165 = lean_ctor_get(x_159, 0); lean_inc(x_165); -lean_dec(x_157); -x_166 = lean_ctor_get(x_164, 0); +lean_dec(x_159); +x_166 = lean_ctor_get(x_158, 1); lean_inc(x_166); -x_167 = lean_ctor_get(x_164, 1); +lean_dec(x_158); +x_167 = lean_ctor_get(x_165, 0); lean_inc(x_167); -lean_dec(x_164); -x_168 = l_Lean_LocalDecl_userName(x_39); +x_168 = lean_ctor_get(x_165, 1); +lean_inc(x_168); +lean_dec(x_165); +x_169 = l_Lean_LocalDecl_userName(x_39); lean_dec(x_39); -x_169 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_169, 0, x_168); -lean_ctor_set(x_169, 1, x_167); -lean_ctor_set(x_169, 2, x_166); -x_170 = lean_array_push(x_36, x_169); -lean_ctor_set(x_28, 1, x_154); -lean_ctor_set(x_28, 0, x_170); +x_170 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_168); +lean_ctor_set(x_170, 2, x_167); +x_171 = lean_array_push(x_37, x_170); +lean_ctor_set(x_28, 1, x_171); +lean_ctor_set(x_28, 0, x_155); lean_inc(x_4); -x_171 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_171, 0, x_4); -lean_ctor_set(x_171, 1, x_26); -x_172 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_172, 0, x_171); -x_17 = x_172; -x_18 = x_165; +x_172 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_172, 0, x_4); +lean_ctor_set(x_172, 1, x_26); +x_173 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_173, 0, x_172); +x_17 = x_173; +x_18 = x_166; goto block_25; } } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; -lean_dec(x_154); -lean_free_object(x_58); +lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +lean_dec(x_155); +lean_free_object(x_59); lean_dec(x_39); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38146,131 +38178,131 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_173 = lean_ctor_get(x_157, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_157, 1); +x_174 = lean_ctor_get(x_158, 0); lean_inc(x_174); -if (lean_is_exclusive(x_157)) { - lean_ctor_release(x_157, 0); - lean_ctor_release(x_157, 1); - x_175 = x_157; +x_175 = lean_ctor_get(x_158, 1); +lean_inc(x_175); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_176 = x_158; } else { - lean_dec_ref(x_157); - x_175 = lean_box(0); + lean_dec_ref(x_158); + x_176 = lean_box(0); } -if (lean_is_scalar(x_175)) { - x_176 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_176)) { + x_177 = lean_alloc_ctor(1, 2, 0); } else { - x_176 = x_175; + x_177 = x_176; } -lean_ctor_set(x_176, 0, x_173); -lean_ctor_set(x_176, 1, x_174); -return 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_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; uint8_t x_181; lean_object* x_182; -lean_dec(x_58); -x_177 = lean_ctor_get(x_55, 1); -lean_inc(x_177); -lean_dec(x_55); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; +lean_dec(x_59); x_178 = lean_ctor_get(x_56, 1); lean_inc(x_178); -if (lean_is_exclusive(x_56)) { - lean_ctor_release(x_56, 0); - lean_ctor_release(x_56, 1); - x_179 = x_56; +lean_dec(x_56); +x_179 = lean_ctor_get(x_57, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_180 = x_57; } else { - lean_dec_ref(x_56); - x_179 = lean_box(0); + lean_dec_ref(x_57); + x_180 = lean_box(0); } -x_180 = l_Lean_Expr_fvar___override(x_16); -x_181 = 1; +x_181 = l_Lean_Expr_fvar___override(x_16); +x_182 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_30); -x_182 = l_Lean_Meta_applySimpResultToProp(x_30, x_180, x_43, x_57, x_181, x_9, x_10, x_11, x_12, x_177); +x_183 = l_Lean_Meta_applySimpResultToProp(x_30, x_181, x_43, x_58, x_182, x_9, x_10, x_11, x_12, x_178); lean_dec(x_43); -if (lean_obj_tag(x_182) == 0) -{ -lean_object* x_183; -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); if (lean_obj_tag(x_183) == 0) { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; -lean_dec(x_39); -x_184 = lean_ctor_get(x_182, 1); +lean_object* x_184; +x_184 = lean_ctor_get(x_183, 0); lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_box(0); -lean_inc(x_178); -if (lean_is_scalar(x_179)) { - x_186 = lean_alloc_ctor(0, 2, 0); +if (lean_obj_tag(x_184) == 0) +{ +lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; +lean_dec(x_39); +x_185 = lean_ctor_get(x_183, 1); +lean_inc(x_185); +lean_dec(x_183); +x_186 = lean_box(0); +lean_inc(x_179); +if (lean_is_scalar(x_180)) { + x_187 = lean_alloc_ctor(0, 2, 0); } else { - x_186 = x_179; + x_187 = x_180; } -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_178); -x_187 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_28, 1, x_178); -x_188 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_187, 1, x_179); +x_188 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_26); -x_189 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_28, 0, x_179); +x_189 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_189, 0, x_188); -x_17 = x_189; -x_18 = x_184; +lean_ctor_set(x_189, 1, x_26); +x_190 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_190, 0, x_189); +x_17 = x_190; +x_18 = x_185; goto block_25; } else { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; -lean_dec(x_179); -x_190 = lean_ctor_get(x_183, 0); -lean_inc(x_190); -lean_dec(x_183); -x_191 = lean_ctor_get(x_182, 1); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_180); +x_191 = lean_ctor_get(x_184, 0); lean_inc(x_191); -lean_dec(x_182); -x_192 = lean_ctor_get(x_190, 0); +lean_dec(x_184); +x_192 = lean_ctor_get(x_183, 1); lean_inc(x_192); -x_193 = lean_ctor_get(x_190, 1); +lean_dec(x_183); +x_193 = lean_ctor_get(x_191, 0); lean_inc(x_193); -lean_dec(x_190); -x_194 = l_Lean_LocalDecl_userName(x_39); +x_194 = lean_ctor_get(x_191, 1); +lean_inc(x_194); +lean_dec(x_191); +x_195 = l_Lean_LocalDecl_userName(x_39); lean_dec(x_39); -x_195 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_195, 0, x_194); -lean_ctor_set(x_195, 1, x_193); -lean_ctor_set(x_195, 2, x_192); -x_196 = lean_array_push(x_36, x_195); -lean_ctor_set(x_28, 1, x_178); -lean_ctor_set(x_28, 0, x_196); +x_196 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_196, 0, x_195); +lean_ctor_set(x_196, 1, x_194); +lean_ctor_set(x_196, 2, x_193); +x_197 = lean_array_push(x_37, x_196); +lean_ctor_set(x_28, 1, x_197); +lean_ctor_set(x_28, 0, x_179); lean_inc(x_4); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_4); -lean_ctor_set(x_197, 1, x_26); -x_198 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_198, 0, x_197); -x_17 = x_198; -x_18 = x_191; +x_198 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_198, 0, x_4); +lean_ctor_set(x_198, 1, x_26); +x_199 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_199, 0, x_198); +x_17 = x_199; +x_18 = x_192; goto block_25; } } else { -lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_dec(x_180); lean_dec(x_179); -lean_dec(x_178); lean_dec(x_39); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38283,37 +38315,37 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_199 = lean_ctor_get(x_182, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_182, 1); +x_200 = lean_ctor_get(x_183, 0); lean_inc(x_200); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_201 = x_182; +x_201 = lean_ctor_get(x_183, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_183)) { + lean_ctor_release(x_183, 0); + lean_ctor_release(x_183, 1); + x_202 = x_183; } else { - lean_dec_ref(x_182); - x_201 = lean_box(0); + lean_dec_ref(x_183); + x_202 = lean_box(0); } -if (lean_is_scalar(x_201)) { - x_202 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_202)) { + x_203 = lean_alloc_ctor(1, 2, 0); } else { - x_202 = x_201; + x_203 = x_202; } -lean_ctor_set(x_202, 0, x_199); -lean_ctor_set(x_202, 1, x_200); -return x_202; +lean_ctor_set(x_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; } } } } else { -uint8_t x_203; +uint8_t x_204; lean_dec(x_43); lean_dec(x_39); lean_free_object(x_28); -lean_dec(x_36); +lean_dec(x_37); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38327,29 +38359,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_203 = !lean_is_exclusive(x_55); -if (x_203 == 0) +x_204 = !lean_is_exclusive(x_56); +if (x_204 == 0) { -return x_55; +return x_56; } else { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_55, 0); -x_205 = lean_ctor_get(x_55, 1); +lean_object* x_205; lean_object* x_206; lean_object* x_207; +x_205 = lean_ctor_get(x_56, 0); +x_206 = lean_ctor_get(x_56, 1); +lean_inc(x_206); lean_inc(x_205); -lean_inc(x_204); -lean_dec(x_55); -x_206 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -return x_206; +lean_dec(x_56); +x_207 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +return x_207; } } } else { -uint8_t x_207; +uint8_t x_208; lean_free_object(x_28); lean_dec(x_37); lean_dec(x_36); @@ -38366,139 +38398,142 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_207 = !lean_is_exclusive(x_38); -if (x_207 == 0) +x_208 = !lean_is_exclusive(x_38); +if (x_208 == 0) { return x_38; } else { -lean_object* x_208; lean_object* x_209; lean_object* x_210; -x_208 = lean_ctor_get(x_38, 0); -x_209 = lean_ctor_get(x_38, 1); +lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_209 = lean_ctor_get(x_38, 0); +x_210 = lean_ctor_get(x_38, 1); +lean_inc(x_210); lean_inc(x_209); -lean_inc(x_208); lean_dec(x_38); -x_210 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_210, 0, x_208); -lean_ctor_set(x_210, 1, x_209); -return x_210; +x_211 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +return x_211; } } } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_28, 0); -x_212 = lean_ctor_get(x_28, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_28, 0); +x_213 = lean_ctor_get(x_28, 1); +lean_inc(x_213); lean_inc(x_212); -lean_inc(x_211); lean_dec(x_28); lean_inc(x_9); lean_inc(x_16); -x_213 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_213) == 0) +x_214 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_214) == 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; uint32_t x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint32_t x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -x_214 = lean_ctor_get(x_213, 0); -lean_inc(x_214); -x_215 = lean_ctor_get(x_213, 1); +lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; uint32_t x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint32_t x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; +x_215 = lean_ctor_get(x_214, 0); lean_inc(x_215); -lean_dec(x_213); -x_216 = l_Lean_LocalDecl_type(x_214); -x_217 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_216, x_9, x_10, x_11, x_12, x_215); -x_218 = lean_ctor_get(x_217, 0); -lean_inc(x_218); -x_219 = lean_ctor_get(x_217, 1); +x_216 = lean_ctor_get(x_214, 1); +lean_inc(x_216); +lean_dec(x_214); +x_217 = l_Lean_LocalDecl_type(x_215); +x_218 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_217, x_9, x_10, x_11, x_12, x_216); +x_219 = lean_ctor_get(x_218, 0); lean_inc(x_219); -lean_dec(x_217); -x_220 = lean_ctor_get(x_1, 0); +x_220 = lean_ctor_get(x_218, 1); lean_inc(x_220); -x_221 = lean_ctor_get_uint32(x_1, sizeof(void*)*4); -x_222 = lean_ctor_get(x_1, 1); -lean_inc(x_222); -x_223 = lean_ctor_get(x_1, 2); +lean_dec(x_218); +x_221 = lean_ctor_get(x_1, 0); +lean_inc(x_221); +x_222 = lean_ctor_get_uint32(x_1, sizeof(void*)*5); +x_223 = lean_ctor_get(x_1, 1); lean_inc(x_223); -x_224 = lean_ctor_get(x_1, 3); +x_224 = lean_ctor_get(x_1, 2); lean_inc(x_224); -x_225 = lean_ctor_get_uint32(x_1, sizeof(void*)*4 + 4); -x_226 = l_Lean_LocalDecl_fvarId(x_214); -x_227 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_227, 0, x_226); -x_228 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_222, x_227); -x_229 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_229, 0, x_220); -lean_ctor_set(x_229, 1, x_228); -lean_ctor_set(x_229, 2, x_223); -lean_ctor_set(x_229, 3, x_224); -lean_ctor_set_uint32(x_229, sizeof(void*)*4, x_221); -lean_ctor_set_uint32(x_229, sizeof(void*)*4 + 4, x_225); +x_225 = lean_ctor_get(x_1, 3); +lean_inc(x_225); +x_226 = lean_ctor_get_uint32(x_1, sizeof(void*)*5 + 4); +x_227 = lean_ctor_get(x_1, 4); +lean_inc(x_227); +x_228 = l_Lean_LocalDecl_fvarId(x_215); +x_229 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_229, 0, x_228); +x_230 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_223, x_229); +x_231 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_231, 0, x_221); +lean_ctor_set(x_231, 1, x_230); +lean_ctor_set(x_231, 2, x_224); +lean_ctor_set(x_231, 3, x_225); +lean_ctor_set(x_231, 4, x_227); +lean_ctor_set_uint32(x_231, sizeof(void*)*5, x_222); +lean_ctor_set_uint32(x_231, sizeof(void*)*5 + 4, x_226); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_218); -x_230 = l_Lean_Meta_simp(x_218, x_229, x_2, x_3, x_212, x_9, x_10, x_11, x_12, x_219); -if (lean_obj_tag(x_230) == 0) +lean_inc(x_219); +x_232 = l_Lean_Meta_simp(x_219, x_231, x_2, x_3, x_212, x_9, x_10, x_11, x_12, x_220); +if (lean_obj_tag(x_232) == 0) { -lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_ctor_get(x_231, 0); -lean_inc(x_232); -x_233 = lean_ctor_get(x_232, 1); +lean_object* x_233; lean_object* x_234; lean_object* x_235; +x_233 = lean_ctor_get(x_232, 0); lean_inc(x_233); -if (lean_obj_tag(x_233) == 0) -{ -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; uint8_t x_238; -lean_dec(x_218); -lean_dec(x_214); -x_234 = lean_ctor_get(x_230, 1); +x_234 = lean_ctor_get(x_233, 0); lean_inc(x_234); -lean_dec(x_230); -x_235 = lean_ctor_get(x_231, 1); +x_235 = lean_ctor_get(x_234, 1); lean_inc(x_235); -if (lean_is_exclusive(x_231)) { - lean_ctor_release(x_231, 0); - lean_ctor_release(x_231, 1); - x_236 = x_231; +if (lean_obj_tag(x_235) == 0) +{ +lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; uint8_t x_240; +lean_dec(x_219); +lean_dec(x_215); +x_236 = lean_ctor_get(x_232, 1); +lean_inc(x_236); +lean_dec(x_232); +x_237 = lean_ctor_get(x_233, 1); +lean_inc(x_237); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_238 = x_233; } else { - lean_dec_ref(x_231); - x_236 = lean_box(0); + lean_dec_ref(x_233); + x_238 = lean_box(0); } -x_237 = lean_ctor_get(x_232, 0); -lean_inc(x_237); -x_238 = l_Lean_Expr_isFalse(x_237); -if (x_238 == 0) +x_239 = lean_ctor_get(x_234, 0); +lean_inc(x_239); +x_240 = l_Lean_Expr_isFalse(x_239); +if (x_240 == 0) { -lean_object* x_239; lean_object* x_240; -lean_dec(x_236); +lean_object* x_241; lean_object* x_242; +lean_dec(x_238); lean_free_object(x_27); lean_free_object(x_26); -x_239 = lean_box(0); +x_241 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_4); -x_240 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_232, x_16, x_211, x_235, x_4, x_30, x_33, x_239, x_9, x_10, x_11, x_12, x_234); -if (lean_obj_tag(x_240) == 0) +x_242 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_234, x_16, x_237, x_213, x_4, x_30, x_33, x_241, x_9, x_10, x_11, x_12, x_236); +if (lean_obj_tag(x_242) == 0) { -lean_object* x_241; lean_object* x_242; -x_241 = lean_ctor_get(x_240, 0); -lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); -lean_dec(x_240); -x_17 = x_241; -x_18 = x_242; +lean_object* x_243; lean_object* x_244; +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); +lean_inc(x_244); +lean_dec(x_242); +x_17 = x_243; +x_18 = x_244; goto block_25; } else { -lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; +lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -38507,91 +38542,91 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_243 = lean_ctor_get(x_240, 0); -lean_inc(x_243); -x_244 = lean_ctor_get(x_240, 1); -lean_inc(x_244); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_245 = x_240; +x_245 = lean_ctor_get(x_242, 0); +lean_inc(x_245); +x_246 = lean_ctor_get(x_242, 1); +lean_inc(x_246); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_247 = x_242; } else { - lean_dec_ref(x_240); - x_245 = lean_box(0); + lean_dec_ref(x_242); + x_247 = lean_box(0); } -if (lean_is_scalar(x_245)) { - x_246 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_247)) { + x_248 = lean_alloc_ctor(1, 2, 0); } else { - x_246 = x_245; + x_248 = x_247; } -lean_ctor_set(x_246, 0, x_243); -lean_ctor_set(x_246, 1, x_244); -return x_246; +lean_ctor_set(x_248, 0, x_245); +lean_ctor_set(x_248, 1, x_246); +return x_248; } } else { -lean_object* x_247; -lean_dec(x_232); +lean_object* x_249; +lean_dec(x_234); lean_inc(x_30); -x_247 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_234); -if (lean_obj_tag(x_247) == 0) +x_249 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_236); +if (lean_obj_tag(x_249) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -lean_dec(x_247); -x_250 = l_Lean_Expr_fvar___override(x_16); +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); +lean_inc(x_251); +lean_dec(x_249); +x_252 = l_Lean_Expr_fvar___override(x_16); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_251 = l_Lean_Meta_mkFalseElim(x_248, x_250, x_9, x_10, x_11, x_12, x_249); -if (lean_obj_tag(x_251) == 0) +x_253 = l_Lean_Meta_mkFalseElim(x_250, x_252, x_9, x_10, x_11, x_12, x_251); +if (lean_obj_tag(x_253) == 0) { -lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_252 = lean_ctor_get(x_251, 0); -lean_inc(x_252); -x_253 = lean_ctor_get(x_251, 1); -lean_inc(x_253); -lean_dec(x_251); -lean_inc(x_30); -x_254 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_252, x_9, x_10, x_11, x_12, x_253); -x_255 = lean_ctor_get(x_254, 1); +lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; +x_254 = lean_ctor_get(x_253, 0); +lean_inc(x_254); +x_255 = lean_ctor_get(x_253, 1); lean_inc(x_255); -lean_dec(x_254); -x_256 = lean_box(0); -lean_inc(x_235); -if (lean_is_scalar(x_236)) { - x_257 = lean_alloc_ctor(0, 2, 0); -} else { - x_257 = x_236; -} -lean_ctor_set(x_257, 0, x_256); -lean_ctor_set(x_257, 1, x_235); -x_258 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_258, 0, x_257); -x_259 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_259, 0, x_211); -lean_ctor_set(x_259, 1, x_235); -lean_ctor_set(x_27, 1, x_259); -x_260 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_260, 0, x_258); -lean_ctor_set(x_260, 1, x_26); -x_261 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_261, 0, x_260); -x_17 = x_261; -x_18 = x_255; +lean_dec(x_253); +lean_inc(x_30); +x_256 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_254, x_9, x_10, x_11, x_12, x_255); +x_257 = lean_ctor_get(x_256, 1); +lean_inc(x_257); +lean_dec(x_256); +x_258 = lean_box(0); +lean_inc(x_237); +if (lean_is_scalar(x_238)) { + x_259 = lean_alloc_ctor(0, 2, 0); +} else { + x_259 = x_238; +} +lean_ctor_set(x_259, 0, x_258); +lean_ctor_set(x_259, 1, x_237); +x_260 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_260, 0, x_259); +x_261 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_261, 0, x_237); +lean_ctor_set(x_261, 1, x_213); +lean_ctor_set(x_27, 1, x_261); +x_262 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_262, 0, x_260); +lean_ctor_set(x_262, 1, x_26); +x_263 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_263, 0, x_262); +x_17 = x_263; +x_18 = x_257; goto block_25; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; -lean_dec(x_236); -lean_dec(x_235); -lean_dec(x_211); +lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; +lean_dec(x_238); +lean_dec(x_237); +lean_dec(x_213); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38604,34 +38639,34 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_262 = lean_ctor_get(x_251, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_251, 1); -lean_inc(x_263); -if (lean_is_exclusive(x_251)) { - lean_ctor_release(x_251, 0); - lean_ctor_release(x_251, 1); - x_264 = x_251; +x_264 = lean_ctor_get(x_253, 0); +lean_inc(x_264); +x_265 = lean_ctor_get(x_253, 1); +lean_inc(x_265); +if (lean_is_exclusive(x_253)) { + lean_ctor_release(x_253, 0); + lean_ctor_release(x_253, 1); + x_266 = x_253; } else { - lean_dec_ref(x_251); - x_264 = lean_box(0); + lean_dec_ref(x_253); + x_266 = lean_box(0); } -if (lean_is_scalar(x_264)) { - x_265 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_266)) { + x_267 = lean_alloc_ctor(1, 2, 0); } else { - x_265 = x_264; + x_267 = x_266; } -lean_ctor_set(x_265, 0, x_262); -lean_ctor_set(x_265, 1, x_263); -return x_265; +lean_ctor_set(x_267, 0, x_264); +lean_ctor_set(x_267, 1, x_265); +return x_267; } } else -{ -lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; -lean_dec(x_236); -lean_dec(x_235); -lean_dec(x_211); +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +lean_dec(x_238); +lean_dec(x_237); +lean_dec(x_213); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38645,147 +38680,147 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_266 = lean_ctor_get(x_247, 0); -lean_inc(x_266); -x_267 = lean_ctor_get(x_247, 1); -lean_inc(x_267); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_268 = x_247; +x_268 = lean_ctor_get(x_249, 0); +lean_inc(x_268); +x_269 = lean_ctor_get(x_249, 1); +lean_inc(x_269); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_270 = x_249; } else { - lean_dec_ref(x_247); - x_268 = lean_box(0); + lean_dec_ref(x_249); + x_270 = lean_box(0); } -if (lean_is_scalar(x_268)) { - x_269 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_270)) { + x_271 = lean_alloc_ctor(1, 2, 0); } else { - x_269 = x_268; + x_271 = x_270; } -lean_ctor_set(x_269, 0, x_266); -lean_ctor_set(x_269, 1, x_267); -return x_269; +lean_ctor_set(x_271, 0, x_268); +lean_ctor_set(x_271, 1, x_269); +return x_271; } } } else { -lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; uint8_t x_275; lean_object* x_276; +lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; uint8_t x_277; lean_object* x_278; +if (lean_is_exclusive(x_235)) { + lean_ctor_release(x_235, 0); + x_272 = x_235; +} else { + lean_dec_ref(x_235); + x_272 = lean_box(0); +} +x_273 = lean_ctor_get(x_232, 1); +lean_inc(x_273); +lean_dec(x_232); +x_274 = lean_ctor_get(x_233, 1); +lean_inc(x_274); if (lean_is_exclusive(x_233)) { lean_ctor_release(x_233, 0); - x_270 = x_233; + lean_ctor_release(x_233, 1); + x_275 = x_233; } else { lean_dec_ref(x_233); - x_270 = lean_box(0); -} -x_271 = lean_ctor_get(x_230, 1); -lean_inc(x_271); -lean_dec(x_230); -x_272 = lean_ctor_get(x_231, 1); -lean_inc(x_272); -if (lean_is_exclusive(x_231)) { - lean_ctor_release(x_231, 0); - lean_ctor_release(x_231, 1); - x_273 = x_231; -} else { - lean_dec_ref(x_231); - x_273 = lean_box(0); + x_275 = lean_box(0); } -x_274 = l_Lean_Expr_fvar___override(x_16); -x_275 = 1; +x_276 = l_Lean_Expr_fvar___override(x_16); +x_277 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_30); -x_276 = l_Lean_Meta_applySimpResultToProp(x_30, x_274, x_218, x_232, x_275, x_9, x_10, x_11, x_12, x_271); -lean_dec(x_218); -if (lean_obj_tag(x_276) == 0) +x_278 = l_Lean_Meta_applySimpResultToProp(x_30, x_276, x_219, x_234, x_277, x_9, x_10, x_11, x_12, x_273); +lean_dec(x_219); +if (lean_obj_tag(x_278) == 0) { -lean_object* x_277; -x_277 = lean_ctor_get(x_276, 0); -lean_inc(x_277); -if (lean_obj_tag(x_277) == 0) +lean_object* x_279; +x_279 = lean_ctor_get(x_278, 0); +lean_inc(x_279); +if (lean_obj_tag(x_279) == 0) { -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -lean_dec(x_214); -x_278 = lean_ctor_get(x_276, 1); -lean_inc(x_278); -lean_dec(x_276); -x_279 = lean_box(0); -lean_inc(x_272); -if (lean_is_scalar(x_273)) { - x_280 = lean_alloc_ctor(0, 2, 0); +lean_object* x_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_dec(x_215); +x_280 = lean_ctor_get(x_278, 1); +lean_inc(x_280); +lean_dec(x_278); +x_281 = lean_box(0); +lean_inc(x_274); +if (lean_is_scalar(x_275)) { + x_282 = lean_alloc_ctor(0, 2, 0); } else { - x_280 = x_273; + x_282 = x_275; } -lean_ctor_set(x_280, 0, x_279); -lean_ctor_set(x_280, 1, x_272); -if (lean_is_scalar(x_270)) { - x_281 = lean_alloc_ctor(1, 1, 0); -} else { - x_281 = x_270; -} -lean_ctor_set(x_281, 0, x_280); -x_282 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_282, 0, x_211); -lean_ctor_set(x_282, 1, x_272); -lean_ctor_set(x_27, 1, x_282); -x_283 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_283, 0, x_281); -lean_ctor_set(x_283, 1, x_26); -x_284 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_284, 0, x_283); -x_17 = x_284; -x_18 = x_278; +lean_ctor_set(x_282, 0, x_281); +lean_ctor_set(x_282, 1, x_274); +if (lean_is_scalar(x_272)) { + x_283 = lean_alloc_ctor(1, 1, 0); +} else { + x_283 = x_272; +} +lean_ctor_set(x_283, 0, x_282); +x_284 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_284, 0, x_274); +lean_ctor_set(x_284, 1, x_213); +lean_ctor_set(x_27, 1, x_284); +x_285 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_285, 0, x_283); +lean_ctor_set(x_285, 1, x_26); +x_286 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_286, 0, x_285); +x_17 = x_286; +x_18 = x_280; goto block_25; } else { -lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; -lean_dec(x_273); -lean_dec(x_270); -x_285 = lean_ctor_get(x_277, 0); -lean_inc(x_285); -lean_dec(x_277); -x_286 = lean_ctor_get(x_276, 1); -lean_inc(x_286); -lean_dec(x_276); -x_287 = lean_ctor_get(x_285, 0); +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec(x_275); +lean_dec(x_272); +x_287 = lean_ctor_get(x_279, 0); lean_inc(x_287); -x_288 = lean_ctor_get(x_285, 1); +lean_dec(x_279); +x_288 = lean_ctor_get(x_278, 1); lean_inc(x_288); -lean_dec(x_285); -x_289 = l_Lean_LocalDecl_userName(x_214); -lean_dec(x_214); -x_290 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_290, 0, x_289); -lean_ctor_set(x_290, 1, x_288); -lean_ctor_set(x_290, 2, x_287); -x_291 = lean_array_push(x_211, x_290); -x_292 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_278); +x_289 = lean_ctor_get(x_287, 0); +lean_inc(x_289); +x_290 = lean_ctor_get(x_287, 1); +lean_inc(x_290); +lean_dec(x_287); +x_291 = l_Lean_LocalDecl_userName(x_215); +lean_dec(x_215); +x_292 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_292, 0, x_291); -lean_ctor_set(x_292, 1, x_272); -lean_ctor_set(x_27, 1, x_292); +lean_ctor_set(x_292, 1, x_290); +lean_ctor_set(x_292, 2, x_289); +x_293 = lean_array_push(x_213, x_292); +x_294 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_294, 0, x_274); +lean_ctor_set(x_294, 1, x_293); +lean_ctor_set(x_27, 1, x_294); lean_inc(x_4); -x_293 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_293, 0, x_4); -lean_ctor_set(x_293, 1, x_26); -x_294 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_294, 0, x_293); -x_17 = x_294; -x_18 = x_286; +x_295 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_295, 0, x_4); +lean_ctor_set(x_295, 1, x_26); +x_296 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_296, 0, x_295); +x_17 = x_296; +x_18 = x_288; goto block_25; } } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; lean_object* x_298; -lean_dec(x_273); +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_275); +lean_dec(x_274); lean_dec(x_272); -lean_dec(x_270); -lean_dec(x_214); -lean_dec(x_211); +lean_dec(x_215); +lean_dec(x_213); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38798,35 +38833,35 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_295 = lean_ctor_get(x_276, 0); -lean_inc(x_295); -x_296 = lean_ctor_get(x_276, 1); -lean_inc(x_296); -if (lean_is_exclusive(x_276)) { - lean_ctor_release(x_276, 0); - lean_ctor_release(x_276, 1); - x_297 = x_276; +x_297 = lean_ctor_get(x_278, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_278, 1); +lean_inc(x_298); +if (lean_is_exclusive(x_278)) { + lean_ctor_release(x_278, 0); + lean_ctor_release(x_278, 1); + x_299 = x_278; } else { - lean_dec_ref(x_276); - x_297 = lean_box(0); + lean_dec_ref(x_278); + x_299 = lean_box(0); } -if (lean_is_scalar(x_297)) { - x_298 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_299)) { + x_300 = lean_alloc_ctor(1, 2, 0); } else { - x_298 = x_297; + x_300 = x_299; } -lean_ctor_set(x_298, 0, x_295); -lean_ctor_set(x_298, 1, x_296); -return x_298; +lean_ctor_set(x_300, 0, x_297); +lean_ctor_set(x_300, 1, x_298); +return x_300; } } } else { -lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; -lean_dec(x_218); -lean_dec(x_214); -lean_dec(x_211); +lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; +lean_dec(x_219); +lean_dec(x_215); +lean_dec(x_213); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38840,33 +38875,33 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_299 = lean_ctor_get(x_230, 0); -lean_inc(x_299); -x_300 = lean_ctor_get(x_230, 1); -lean_inc(x_300); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_301 = x_230; +x_301 = lean_ctor_get(x_232, 0); +lean_inc(x_301); +x_302 = lean_ctor_get(x_232, 1); +lean_inc(x_302); +if (lean_is_exclusive(x_232)) { + lean_ctor_release(x_232, 0); + lean_ctor_release(x_232, 1); + x_303 = x_232; } else { - lean_dec_ref(x_230); - x_301 = lean_box(0); + lean_dec_ref(x_232); + x_303 = lean_box(0); } -if (lean_is_scalar(x_301)) { - x_302 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_303)) { + x_304 = lean_alloc_ctor(1, 2, 0); } else { - x_302 = x_301; + x_304 = x_303; } -lean_ctor_set(x_302, 0, x_299); -lean_ctor_set(x_302, 1, x_300); -return x_302; +lean_ctor_set(x_304, 0, x_301); +lean_ctor_set(x_304, 1, x_302); +return x_304; } } else { -lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; +lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; +lean_dec(x_213); lean_dec(x_212); -lean_dec(x_211); lean_free_object(x_27); lean_dec(x_33); lean_free_object(x_26); @@ -38880,152 +38915,155 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_303 = lean_ctor_get(x_213, 0); -lean_inc(x_303); -x_304 = lean_ctor_get(x_213, 1); -lean_inc(x_304); -if (lean_is_exclusive(x_213)) { - lean_ctor_release(x_213, 0); - lean_ctor_release(x_213, 1); - x_305 = x_213; +x_305 = lean_ctor_get(x_214, 0); +lean_inc(x_305); +x_306 = lean_ctor_get(x_214, 1); +lean_inc(x_306); +if (lean_is_exclusive(x_214)) { + lean_ctor_release(x_214, 0); + lean_ctor_release(x_214, 1); + x_307 = x_214; } else { - lean_dec_ref(x_213); - x_305 = lean_box(0); + lean_dec_ref(x_214); + x_307 = lean_box(0); } -if (lean_is_scalar(x_305)) { - x_306 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_307)) { + x_308 = lean_alloc_ctor(1, 2, 0); } else { - x_306 = x_305; + x_308 = x_307; } -lean_ctor_set(x_306, 0, x_303); -lean_ctor_set(x_306, 1, x_304); -return x_306; +lean_ctor_set(x_308, 0, x_305); +lean_ctor_set(x_308, 1, x_306); +return x_308; } } } else { -lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -x_307 = lean_ctor_get(x_27, 0); -lean_inc(x_307); -lean_dec(x_27); -x_308 = lean_ctor_get(x_28, 0); -lean_inc(x_308); -x_309 = lean_ctor_get(x_28, 1); +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; lean_object* x_313; +x_309 = lean_ctor_get(x_27, 0); lean_inc(x_309); +lean_dec(x_27); +x_310 = lean_ctor_get(x_28, 0); +lean_inc(x_310); +x_311 = lean_ctor_get(x_28, 1); +lean_inc(x_311); if (lean_is_exclusive(x_28)) { lean_ctor_release(x_28, 0); lean_ctor_release(x_28, 1); - x_310 = x_28; + x_312 = x_28; } else { lean_dec_ref(x_28); - x_310 = lean_box(0); + x_312 = lean_box(0); } lean_inc(x_9); lean_inc(x_16); -x_311 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_311) == 0) +x_313 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_313) == 0) { -lean_object* x_312; lean_object* x_313; lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; uint32_t x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; uint32_t x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; -x_312 = lean_ctor_get(x_311, 0); -lean_inc(x_312); -x_313 = lean_ctor_get(x_311, 1); -lean_inc(x_313); -lean_dec(x_311); -x_314 = l_Lean_LocalDecl_type(x_312); -x_315 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_314, x_9, x_10, x_11, x_12, x_313); -x_316 = lean_ctor_get(x_315, 0); -lean_inc(x_316); -x_317 = lean_ctor_get(x_315, 1); -lean_inc(x_317); -lean_dec(x_315); -x_318 = lean_ctor_get(x_1, 0); +lean_object* x_314; lean_object* x_315; lean_object* x_316; lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; uint32_t x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; uint32_t x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; lean_object* x_331; +x_314 = lean_ctor_get(x_313, 0); +lean_inc(x_314); +x_315 = lean_ctor_get(x_313, 1); +lean_inc(x_315); +lean_dec(x_313); +x_316 = l_Lean_LocalDecl_type(x_314); +x_317 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_316, x_9, x_10, x_11, x_12, x_315); +x_318 = lean_ctor_get(x_317, 0); lean_inc(x_318); -x_319 = lean_ctor_get_uint32(x_1, sizeof(void*)*4); -x_320 = lean_ctor_get(x_1, 1); +x_319 = lean_ctor_get(x_317, 1); +lean_inc(x_319); +lean_dec(x_317); +x_320 = lean_ctor_get(x_1, 0); lean_inc(x_320); -x_321 = lean_ctor_get(x_1, 2); -lean_inc(x_321); -x_322 = lean_ctor_get(x_1, 3); +x_321 = lean_ctor_get_uint32(x_1, sizeof(void*)*5); +x_322 = lean_ctor_get(x_1, 1); lean_inc(x_322); -x_323 = lean_ctor_get_uint32(x_1, sizeof(void*)*4 + 4); -x_324 = l_Lean_LocalDecl_fvarId(x_312); -x_325 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_325, 0, x_324); -x_326 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_320, x_325); -x_327 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_327, 0, x_318); -lean_ctor_set(x_327, 1, x_326); -lean_ctor_set(x_327, 2, x_321); -lean_ctor_set(x_327, 3, x_322); -lean_ctor_set_uint32(x_327, sizeof(void*)*4, x_319); -lean_ctor_set_uint32(x_327, sizeof(void*)*4 + 4, x_323); +x_323 = lean_ctor_get(x_1, 2); +lean_inc(x_323); +x_324 = lean_ctor_get(x_1, 3); +lean_inc(x_324); +x_325 = lean_ctor_get_uint32(x_1, sizeof(void*)*5 + 4); +x_326 = lean_ctor_get(x_1, 4); +lean_inc(x_326); +x_327 = l_Lean_LocalDecl_fvarId(x_314); +x_328 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_328, 0, x_327); +x_329 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_322, x_328); +x_330 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_330, 0, x_320); +lean_ctor_set(x_330, 1, x_329); +lean_ctor_set(x_330, 2, x_323); +lean_ctor_set(x_330, 3, x_324); +lean_ctor_set(x_330, 4, x_326); +lean_ctor_set_uint32(x_330, sizeof(void*)*5, x_321); +lean_ctor_set_uint32(x_330, sizeof(void*)*5 + 4, x_325); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_316); -x_328 = l_Lean_Meta_simp(x_316, x_327, x_2, x_3, x_309, x_9, x_10, x_11, x_12, x_317); -if (lean_obj_tag(x_328) == 0) -{ -lean_object* x_329; lean_object* x_330; lean_object* x_331; -x_329 = lean_ctor_get(x_328, 0); -lean_inc(x_329); -x_330 = lean_ctor_get(x_329, 0); -lean_inc(x_330); -x_331 = lean_ctor_get(x_330, 1); -lean_inc(x_331); +lean_inc(x_318); +x_331 = l_Lean_Meta_simp(x_318, x_330, x_2, x_3, x_310, x_9, x_10, x_11, x_12, x_319); if (lean_obj_tag(x_331) == 0) { -lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; uint8_t x_336; -lean_dec(x_316); -lean_dec(x_312); -x_332 = lean_ctor_get(x_328, 1); +lean_object* x_332; lean_object* x_333; lean_object* x_334; +x_332 = lean_ctor_get(x_331, 0); lean_inc(x_332); -lean_dec(x_328); -x_333 = lean_ctor_get(x_329, 1); +x_333 = lean_ctor_get(x_332, 0); lean_inc(x_333); -if (lean_is_exclusive(x_329)) { - lean_ctor_release(x_329, 0); - lean_ctor_release(x_329, 1); - x_334 = x_329; +x_334 = lean_ctor_get(x_333, 1); +lean_inc(x_334); +if (lean_obj_tag(x_334) == 0) +{ +lean_object* x_335; lean_object* x_336; lean_object* x_337; lean_object* x_338; uint8_t x_339; +lean_dec(x_318); +lean_dec(x_314); +x_335 = lean_ctor_get(x_331, 1); +lean_inc(x_335); +lean_dec(x_331); +x_336 = lean_ctor_get(x_332, 1); +lean_inc(x_336); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_337 = x_332; } else { - lean_dec_ref(x_329); - x_334 = lean_box(0); + lean_dec_ref(x_332); + x_337 = lean_box(0); } -x_335 = lean_ctor_get(x_330, 0); -lean_inc(x_335); -x_336 = l_Lean_Expr_isFalse(x_335); -if (x_336 == 0) +x_338 = lean_ctor_get(x_333, 0); +lean_inc(x_338); +x_339 = l_Lean_Expr_isFalse(x_338); +if (x_339 == 0) { -lean_object* x_337; lean_object* x_338; -lean_dec(x_334); -lean_dec(x_310); +lean_object* x_340; lean_object* x_341; +lean_dec(x_337); +lean_dec(x_312); lean_free_object(x_26); -x_337 = lean_box(0); +x_340 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_4); -x_338 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_330, x_16, x_308, x_333, x_4, x_30, x_307, x_337, x_9, x_10, x_11, x_12, x_332); -if (lean_obj_tag(x_338) == 0) +x_341 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_333, x_16, x_336, x_311, x_4, x_30, x_309, x_340, x_9, x_10, x_11, x_12, x_335); +if (lean_obj_tag(x_341) == 0) { -lean_object* x_339; lean_object* x_340; -x_339 = lean_ctor_get(x_338, 0); -lean_inc(x_339); -x_340 = lean_ctor_get(x_338, 1); -lean_inc(x_340); -lean_dec(x_338); -x_17 = x_339; -x_18 = x_340; +lean_object* x_342; lean_object* x_343; +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_17 = x_342; +x_18 = x_343; goto block_25; } else { -lean_object* x_341; lean_object* x_342; lean_object* x_343; lean_object* x_344; +lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -39034,100 +39072,100 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_341 = lean_ctor_get(x_338, 0); -lean_inc(x_341); -x_342 = lean_ctor_get(x_338, 1); -lean_inc(x_342); -if (lean_is_exclusive(x_338)) { - lean_ctor_release(x_338, 0); - lean_ctor_release(x_338, 1); - x_343 = x_338; +x_344 = lean_ctor_get(x_341, 0); +lean_inc(x_344); +x_345 = lean_ctor_get(x_341, 1); +lean_inc(x_345); +if (lean_is_exclusive(x_341)) { + lean_ctor_release(x_341, 0); + lean_ctor_release(x_341, 1); + x_346 = x_341; } else { - lean_dec_ref(x_338); - x_343 = lean_box(0); + lean_dec_ref(x_341); + x_346 = lean_box(0); } -if (lean_is_scalar(x_343)) { - x_344 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_346)) { + x_347 = lean_alloc_ctor(1, 2, 0); } else { - x_344 = x_343; + x_347 = x_346; } -lean_ctor_set(x_344, 0, x_341); -lean_ctor_set(x_344, 1, x_342); -return x_344; +lean_ctor_set(x_347, 0, x_344); +lean_ctor_set(x_347, 1, x_345); +return x_347; } } else { -lean_object* x_345; -lean_dec(x_330); +lean_object* x_348; +lean_dec(x_333); lean_inc(x_30); -x_345 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_332); -if (lean_obj_tag(x_345) == 0) -{ -lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; -x_346 = lean_ctor_get(x_345, 0); -lean_inc(x_346); -x_347 = lean_ctor_get(x_345, 1); -lean_inc(x_347); -lean_dec(x_345); -x_348 = l_Lean_Expr_fvar___override(x_16); +x_348 = l_Lean_MVarId_getType(x_30, x_9, x_10, x_11, x_12, x_335); +if (lean_obj_tag(x_348) == 0) +{ +lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; +x_349 = lean_ctor_get(x_348, 0); +lean_inc(x_349); +x_350 = lean_ctor_get(x_348, 1); +lean_inc(x_350); +lean_dec(x_348); +x_351 = l_Lean_Expr_fvar___override(x_16); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_349 = l_Lean_Meta_mkFalseElim(x_346, x_348, x_9, x_10, x_11, x_12, x_347); -if (lean_obj_tag(x_349) == 0) +x_352 = l_Lean_Meta_mkFalseElim(x_349, x_351, x_9, x_10, x_11, x_12, x_350); +if (lean_obj_tag(x_352) == 0) { -lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; -x_350 = lean_ctor_get(x_349, 0); -lean_inc(x_350); -x_351 = lean_ctor_get(x_349, 1); -lean_inc(x_351); -lean_dec(x_349); -lean_inc(x_30); -x_352 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_350, x_9, x_10, x_11, x_12, x_351); -x_353 = lean_ctor_get(x_352, 1); +lean_object* x_353; lean_object* x_354; lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; +x_353 = lean_ctor_get(x_352, 0); lean_inc(x_353); +x_354 = lean_ctor_get(x_352, 1); +lean_inc(x_354); lean_dec(x_352); -x_354 = lean_box(0); -lean_inc(x_333); -if (lean_is_scalar(x_334)) { - x_355 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_30); +x_355 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_30, x_353, x_9, x_10, x_11, x_12, x_354); +x_356 = lean_ctor_get(x_355, 1); +lean_inc(x_356); +lean_dec(x_355); +x_357 = lean_box(0); +lean_inc(x_336); +if (lean_is_scalar(x_337)) { + x_358 = lean_alloc_ctor(0, 2, 0); } else { - x_355 = x_334; + x_358 = x_337; } -lean_ctor_set(x_355, 0, x_354); -lean_ctor_set(x_355, 1, x_333); -x_356 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_356, 0, x_355); -if (lean_is_scalar(x_310)) { - x_357 = lean_alloc_ctor(0, 2, 0); -} else { - x_357 = x_310; -} -lean_ctor_set(x_357, 0, x_308); -lean_ctor_set(x_357, 1, x_333); -x_358 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_358, 0, x_307); -lean_ctor_set(x_358, 1, x_357); -lean_ctor_set(x_26, 1, x_358); -x_359 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_359, 0, x_356); -lean_ctor_set(x_359, 1, x_26); -x_360 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_360, 0, x_359); -x_17 = x_360; -x_18 = x_353; +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_336); +x_359 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_359, 0, x_358); +if (lean_is_scalar(x_312)) { + x_360 = lean_alloc_ctor(0, 2, 0); +} else { + x_360 = x_312; +} +lean_ctor_set(x_360, 0, x_336); +lean_ctor_set(x_360, 1, x_311); +x_361 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_361, 0, x_309); +lean_ctor_set(x_361, 1, x_360); +lean_ctor_set(x_26, 1, x_361); +x_362 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_362, 0, x_359); +lean_ctor_set(x_362, 1, x_26); +x_363 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_363, 0, x_362); +x_17 = x_363; +x_18 = x_356; goto block_25; } else { -lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; -lean_dec(x_334); -lean_dec(x_333); -lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); +lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; +lean_dec(x_337); +lean_dec(x_336); +lean_dec(x_312); +lean_dec(x_311); +lean_dec(x_309); lean_free_object(x_26); lean_dec(x_30); lean_dec(x_12); @@ -39138,36 +39176,36 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_361 = lean_ctor_get(x_349, 0); -lean_inc(x_361); -x_362 = lean_ctor_get(x_349, 1); -lean_inc(x_362); -if (lean_is_exclusive(x_349)) { - lean_ctor_release(x_349, 0); - lean_ctor_release(x_349, 1); - x_363 = x_349; +x_364 = lean_ctor_get(x_352, 0); +lean_inc(x_364); +x_365 = lean_ctor_get(x_352, 1); +lean_inc(x_365); +if (lean_is_exclusive(x_352)) { + lean_ctor_release(x_352, 0); + lean_ctor_release(x_352, 1); + x_366 = x_352; } else { - lean_dec_ref(x_349); - x_363 = lean_box(0); + lean_dec_ref(x_352); + x_366 = lean_box(0); } -if (lean_is_scalar(x_363)) { - x_364 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_366)) { + x_367 = lean_alloc_ctor(1, 2, 0); } else { - x_364 = x_363; + x_367 = x_366; } -lean_ctor_set(x_364, 0, x_361); -lean_ctor_set(x_364, 1, x_362); -return x_364; +lean_ctor_set(x_367, 0, x_364); +lean_ctor_set(x_367, 1, x_365); +return x_367; } } else { -lean_object* x_365; lean_object* x_366; lean_object* x_367; lean_object* x_368; -lean_dec(x_334); -lean_dec(x_333); -lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); +lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; +lean_dec(x_337); +lean_dec(x_336); +lean_dec(x_312); +lean_dec(x_311); +lean_dec(x_309); lean_free_object(x_26); lean_dec(x_30); lean_dec(x_16); @@ -39179,163 +39217,163 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_365 = lean_ctor_get(x_345, 0); -lean_inc(x_365); -x_366 = lean_ctor_get(x_345, 1); -lean_inc(x_366); -if (lean_is_exclusive(x_345)) { - lean_ctor_release(x_345, 0); - lean_ctor_release(x_345, 1); - x_367 = x_345; +x_368 = lean_ctor_get(x_348, 0); +lean_inc(x_368); +x_369 = lean_ctor_get(x_348, 1); +lean_inc(x_369); +if (lean_is_exclusive(x_348)) { + lean_ctor_release(x_348, 0); + lean_ctor_release(x_348, 1); + x_370 = x_348; } else { - lean_dec_ref(x_345); - x_367 = lean_box(0); + lean_dec_ref(x_348); + x_370 = lean_box(0); } -if (lean_is_scalar(x_367)) { - x_368 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_370)) { + x_371 = lean_alloc_ctor(1, 2, 0); } else { - x_368 = x_367; + x_371 = x_370; } -lean_ctor_set(x_368, 0, x_365); -lean_ctor_set(x_368, 1, x_366); -return x_368; +lean_ctor_set(x_371, 0, x_368); +lean_ctor_set(x_371, 1, x_369); +return x_371; } } } else { -lean_object* x_369; lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; uint8_t x_374; lean_object* x_375; -if (lean_is_exclusive(x_331)) { - lean_ctor_release(x_331, 0); - x_369 = x_331; +lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; lean_object* x_376; uint8_t x_377; lean_object* x_378; +if (lean_is_exclusive(x_334)) { + lean_ctor_release(x_334, 0); + x_372 = x_334; } else { - lean_dec_ref(x_331); - x_369 = lean_box(0); + lean_dec_ref(x_334); + x_372 = lean_box(0); } -x_370 = lean_ctor_get(x_328, 1); -lean_inc(x_370); -lean_dec(x_328); -x_371 = lean_ctor_get(x_329, 1); -lean_inc(x_371); -if (lean_is_exclusive(x_329)) { - lean_ctor_release(x_329, 0); - lean_ctor_release(x_329, 1); - x_372 = x_329; +x_373 = lean_ctor_get(x_331, 1); +lean_inc(x_373); +lean_dec(x_331); +x_374 = lean_ctor_get(x_332, 1); +lean_inc(x_374); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_375 = x_332; } else { - lean_dec_ref(x_329); - x_372 = lean_box(0); + lean_dec_ref(x_332); + x_375 = lean_box(0); } -x_373 = l_Lean_Expr_fvar___override(x_16); -x_374 = 1; +x_376 = l_Lean_Expr_fvar___override(x_16); +x_377 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_30); -x_375 = l_Lean_Meta_applySimpResultToProp(x_30, x_373, x_316, x_330, x_374, x_9, x_10, x_11, x_12, x_370); -lean_dec(x_316); -if (lean_obj_tag(x_375) == 0) +x_378 = l_Lean_Meta_applySimpResultToProp(x_30, x_376, x_318, x_333, x_377, x_9, x_10, x_11, x_12, x_373); +lean_dec(x_318); +if (lean_obj_tag(x_378) == 0) { -lean_object* x_376; -x_376 = lean_ctor_get(x_375, 0); -lean_inc(x_376); -if (lean_obj_tag(x_376) == 0) +lean_object* x_379; +x_379 = lean_ctor_get(x_378, 0); +lean_inc(x_379); +if (lean_obj_tag(x_379) == 0) { -lean_object* x_377; lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; -lean_dec(x_312); -x_377 = lean_ctor_get(x_375, 1); -lean_inc(x_377); -lean_dec(x_375); -x_378 = lean_box(0); -lean_inc(x_371); +lean_object* x_380; lean_object* x_381; lean_object* x_382; lean_object* x_383; lean_object* x_384; lean_object* x_385; lean_object* x_386; lean_object* x_387; +lean_dec(x_314); +x_380 = lean_ctor_get(x_378, 1); +lean_inc(x_380); +lean_dec(x_378); +x_381 = lean_box(0); +lean_inc(x_374); +if (lean_is_scalar(x_375)) { + x_382 = lean_alloc_ctor(0, 2, 0); +} else { + x_382 = x_375; +} +lean_ctor_set(x_382, 0, x_381); +lean_ctor_set(x_382, 1, x_374); if (lean_is_scalar(x_372)) { - x_379 = lean_alloc_ctor(0, 2, 0); + x_383 = lean_alloc_ctor(1, 1, 0); } else { - x_379 = x_372; + x_383 = x_372; } -lean_ctor_set(x_379, 0, x_378); -lean_ctor_set(x_379, 1, x_371); -if (lean_is_scalar(x_369)) { - x_380 = lean_alloc_ctor(1, 1, 0); -} else { - x_380 = x_369; -} -lean_ctor_set(x_380, 0, x_379); -if (lean_is_scalar(x_310)) { - x_381 = lean_alloc_ctor(0, 2, 0); -} else { - x_381 = x_310; -} -lean_ctor_set(x_381, 0, x_308); -lean_ctor_set(x_381, 1, x_371); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_307); -lean_ctor_set(x_382, 1, x_381); -lean_ctor_set(x_26, 1, x_382); -x_383 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_383, 0, x_380); -lean_ctor_set(x_383, 1, x_26); -x_384 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_384, 0, x_383); -x_17 = x_384; -x_18 = x_377; +lean_ctor_set(x_383, 0, x_382); +if (lean_is_scalar(x_312)) { + x_384 = lean_alloc_ctor(0, 2, 0); +} else { + x_384 = x_312; +} +lean_ctor_set(x_384, 0, x_374); +lean_ctor_set(x_384, 1, x_311); +x_385 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_385, 0, x_309); +lean_ctor_set(x_385, 1, x_384); +lean_ctor_set(x_26, 1, x_385); +x_386 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_386, 0, x_383); +lean_ctor_set(x_386, 1, x_26); +x_387 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_387, 0, x_386); +x_17 = x_387; +x_18 = x_380; goto block_25; } else { -lean_object* x_385; lean_object* x_386; lean_object* x_387; lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; -lean_dec(x_372); -lean_dec(x_369); -x_385 = lean_ctor_get(x_376, 0); -lean_inc(x_385); -lean_dec(x_376); -x_386 = lean_ctor_get(x_375, 1); -lean_inc(x_386); +lean_object* x_388; lean_object* x_389; lean_object* x_390; lean_object* x_391; lean_object* x_392; lean_object* x_393; lean_object* x_394; lean_object* x_395; lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_dec(x_375); -x_387 = lean_ctor_get(x_385, 0); -lean_inc(x_387); -x_388 = lean_ctor_get(x_385, 1); +lean_dec(x_372); +x_388 = lean_ctor_get(x_379, 0); lean_inc(x_388); -lean_dec(x_385); -x_389 = l_Lean_LocalDecl_userName(x_312); -lean_dec(x_312); -x_390 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_390, 0, x_389); -lean_ctor_set(x_390, 1, x_388); -lean_ctor_set(x_390, 2, x_387); -x_391 = lean_array_push(x_308, x_390); -if (lean_is_scalar(x_310)) { - x_392 = lean_alloc_ctor(0, 2, 0); -} else { - x_392 = x_310; -} -lean_ctor_set(x_392, 0, x_391); -lean_ctor_set(x_392, 1, x_371); -x_393 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_393, 0, x_307); -lean_ctor_set(x_393, 1, x_392); -lean_ctor_set(x_26, 1, x_393); +lean_dec(x_379); +x_389 = lean_ctor_get(x_378, 1); +lean_inc(x_389); +lean_dec(x_378); +x_390 = lean_ctor_get(x_388, 0); +lean_inc(x_390); +x_391 = lean_ctor_get(x_388, 1); +lean_inc(x_391); +lean_dec(x_388); +x_392 = l_Lean_LocalDecl_userName(x_314); +lean_dec(x_314); +x_393 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_393, 0, x_392); +lean_ctor_set(x_393, 1, x_391); +lean_ctor_set(x_393, 2, x_390); +x_394 = lean_array_push(x_311, x_393); +if (lean_is_scalar(x_312)) { + x_395 = lean_alloc_ctor(0, 2, 0); +} else { + x_395 = x_312; +} +lean_ctor_set(x_395, 0, x_374); +lean_ctor_set(x_395, 1, x_394); +x_396 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_396, 0, x_309); +lean_ctor_set(x_396, 1, x_395); +lean_ctor_set(x_26, 1, x_396); lean_inc(x_4); -x_394 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_394, 0, x_4); -lean_ctor_set(x_394, 1, x_26); -x_395 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_395, 0, x_394); -x_17 = x_395; -x_18 = x_386; +x_397 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_397, 0, x_4); +lean_ctor_set(x_397, 1, x_26); +x_398 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_398, 0, x_397); +x_17 = x_398; +x_18 = x_389; goto block_25; } } else { -lean_object* x_396; lean_object* x_397; lean_object* x_398; lean_object* x_399; +lean_object* x_399; lean_object* x_400; lean_object* x_401; lean_object* x_402; +lean_dec(x_375); +lean_dec(x_374); lean_dec(x_372); -lean_dec(x_371); -lean_dec(x_369); +lean_dec(x_314); lean_dec(x_312); -lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); +lean_dec(x_311); +lean_dec(x_309); lean_free_object(x_26); lean_dec(x_30); lean_dec(x_12); @@ -39346,37 +39384,37 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_396 = lean_ctor_get(x_375, 0); -lean_inc(x_396); -x_397 = lean_ctor_get(x_375, 1); -lean_inc(x_397); -if (lean_is_exclusive(x_375)) { - lean_ctor_release(x_375, 0); - lean_ctor_release(x_375, 1); - x_398 = x_375; +x_399 = lean_ctor_get(x_378, 0); +lean_inc(x_399); +x_400 = lean_ctor_get(x_378, 1); +lean_inc(x_400); +if (lean_is_exclusive(x_378)) { + lean_ctor_release(x_378, 0); + lean_ctor_release(x_378, 1); + x_401 = x_378; } else { - lean_dec_ref(x_375); - x_398 = lean_box(0); + lean_dec_ref(x_378); + x_401 = lean_box(0); } -if (lean_is_scalar(x_398)) { - x_399 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_401)) { + x_402 = lean_alloc_ctor(1, 2, 0); } else { - x_399 = x_398; + x_402 = x_401; } -lean_ctor_set(x_399, 0, x_396); -lean_ctor_set(x_399, 1, x_397); -return x_399; +lean_ctor_set(x_402, 0, x_399); +lean_ctor_set(x_402, 1, x_400); +return x_402; } } } else { -lean_object* x_400; lean_object* x_401; lean_object* x_402; lean_object* x_403; -lean_dec(x_316); +lean_object* x_403; lean_object* x_404; lean_object* x_405; lean_object* x_406; +lean_dec(x_318); +lean_dec(x_314); lean_dec(x_312); -lean_dec(x_310); -lean_dec(x_308); -lean_dec(x_307); +lean_dec(x_311); +lean_dec(x_309); lean_free_object(x_26); lean_dec(x_30); lean_dec(x_16); @@ -39388,35 +39426,35 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_400 = lean_ctor_get(x_328, 0); -lean_inc(x_400); -x_401 = lean_ctor_get(x_328, 1); -lean_inc(x_401); -if (lean_is_exclusive(x_328)) { - lean_ctor_release(x_328, 0); - lean_ctor_release(x_328, 1); - x_402 = x_328; -} else { - lean_dec_ref(x_328); - x_402 = lean_box(0); +x_403 = lean_ctor_get(x_331, 0); +lean_inc(x_403); +x_404 = lean_ctor_get(x_331, 1); +lean_inc(x_404); +if (lean_is_exclusive(x_331)) { + lean_ctor_release(x_331, 0); + lean_ctor_release(x_331, 1); + x_405 = x_331; +} else { + lean_dec_ref(x_331); + x_405 = lean_box(0); } -if (lean_is_scalar(x_402)) { - x_403 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_405)) { + x_406 = lean_alloc_ctor(1, 2, 0); } else { - x_403 = x_402; + x_406 = x_405; } -lean_ctor_set(x_403, 0, x_400); -lean_ctor_set(x_403, 1, x_401); -return x_403; +lean_ctor_set(x_406, 0, x_403); +lean_ctor_set(x_406, 1, x_404); +return x_406; } } else { -lean_object* x_404; lean_object* x_405; lean_object* x_406; lean_object* x_407; +lean_object* x_407; lean_object* x_408; lean_object* x_409; lean_object* x_410; +lean_dec(x_312); +lean_dec(x_311); lean_dec(x_310); lean_dec(x_309); -lean_dec(x_308); -lean_dec(x_307); lean_free_object(x_26); lean_dec(x_30); lean_dec(x_16); @@ -39428,162 +39466,165 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_404 = lean_ctor_get(x_311, 0); -lean_inc(x_404); -x_405 = lean_ctor_get(x_311, 1); -lean_inc(x_405); -if (lean_is_exclusive(x_311)) { - lean_ctor_release(x_311, 0); - lean_ctor_release(x_311, 1); - x_406 = x_311; +x_407 = lean_ctor_get(x_313, 0); +lean_inc(x_407); +x_408 = lean_ctor_get(x_313, 1); +lean_inc(x_408); +if (lean_is_exclusive(x_313)) { + lean_ctor_release(x_313, 0); + lean_ctor_release(x_313, 1); + x_409 = x_313; } else { - lean_dec_ref(x_311); - x_406 = lean_box(0); + lean_dec_ref(x_313); + x_409 = lean_box(0); } -if (lean_is_scalar(x_406)) { - x_407 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_409)) { + x_410 = lean_alloc_ctor(1, 2, 0); } else { - x_407 = x_406; + x_410 = x_409; } -lean_ctor_set(x_407, 0, x_404); -lean_ctor_set(x_407, 1, x_405); -return x_407; +lean_ctor_set(x_410, 0, x_407); +lean_ctor_set(x_410, 1, x_408); +return x_410; } } } else { -lean_object* x_408; lean_object* x_409; lean_object* x_410; lean_object* x_411; lean_object* x_412; lean_object* x_413; lean_object* x_414; -x_408 = lean_ctor_get(x_26, 0); -lean_inc(x_408); +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; +x_411 = lean_ctor_get(x_26, 0); +lean_inc(x_411); lean_dec(x_26); -x_409 = lean_ctor_get(x_27, 0); -lean_inc(x_409); +x_412 = lean_ctor_get(x_27, 0); +lean_inc(x_412); if (lean_is_exclusive(x_27)) { lean_ctor_release(x_27, 0); lean_ctor_release(x_27, 1); - x_410 = x_27; + x_413 = x_27; } else { lean_dec_ref(x_27); - x_410 = lean_box(0); + x_413 = lean_box(0); } -x_411 = lean_ctor_get(x_28, 0); -lean_inc(x_411); -x_412 = lean_ctor_get(x_28, 1); -lean_inc(x_412); +x_414 = lean_ctor_get(x_28, 0); +lean_inc(x_414); +x_415 = lean_ctor_get(x_28, 1); +lean_inc(x_415); if (lean_is_exclusive(x_28)) { lean_ctor_release(x_28, 0); lean_ctor_release(x_28, 1); - x_413 = x_28; + x_416 = x_28; } else { lean_dec_ref(x_28); - x_413 = lean_box(0); + x_416 = lean_box(0); } lean_inc(x_9); lean_inc(x_16); -x_414 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_414) == 0) +x_417 = l_Lean_FVarId_getDecl(x_16, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_417) == 0) { -lean_object* x_415; lean_object* x_416; lean_object* x_417; lean_object* x_418; lean_object* x_419; lean_object* x_420; lean_object* x_421; uint32_t x_422; lean_object* x_423; lean_object* x_424; lean_object* x_425; uint32_t x_426; lean_object* x_427; lean_object* x_428; lean_object* x_429; lean_object* x_430; lean_object* x_431; -x_415 = lean_ctor_get(x_414, 0); -lean_inc(x_415); -x_416 = lean_ctor_get(x_414, 1); -lean_inc(x_416); -lean_dec(x_414); -x_417 = l_Lean_LocalDecl_type(x_415); -x_418 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_417, x_9, x_10, x_11, x_12, x_416); -x_419 = lean_ctor_get(x_418, 0); +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; uint32_t x_425; lean_object* x_426; lean_object* x_427; lean_object* x_428; uint32_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; +x_418 = lean_ctor_get(x_417, 0); +lean_inc(x_418); +x_419 = lean_ctor_get(x_417, 1); lean_inc(x_419); -x_420 = lean_ctor_get(x_418, 1); -lean_inc(x_420); -lean_dec(x_418); -x_421 = lean_ctor_get(x_1, 0); -lean_inc(x_421); -x_422 = lean_ctor_get_uint32(x_1, sizeof(void*)*4); -x_423 = lean_ctor_get(x_1, 1); +lean_dec(x_417); +x_420 = l_Lean_LocalDecl_type(x_418); +x_421 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_420, x_9, x_10, x_11, x_12, x_419); +x_422 = lean_ctor_get(x_421, 0); +lean_inc(x_422); +x_423 = lean_ctor_get(x_421, 1); lean_inc(x_423); -x_424 = lean_ctor_get(x_1, 2); +lean_dec(x_421); +x_424 = lean_ctor_get(x_1, 0); lean_inc(x_424); -x_425 = lean_ctor_get(x_1, 3); -lean_inc(x_425); -x_426 = lean_ctor_get_uint32(x_1, sizeof(void*)*4 + 4); -x_427 = l_Lean_LocalDecl_fvarId(x_415); -x_428 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_428, 0, x_427); -x_429 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_423, x_428); -x_430 = lean_alloc_ctor(0, 4, 8); -lean_ctor_set(x_430, 0, x_421); -lean_ctor_set(x_430, 1, x_429); -lean_ctor_set(x_430, 2, x_424); -lean_ctor_set(x_430, 3, x_425); -lean_ctor_set_uint32(x_430, sizeof(void*)*4, x_422); -lean_ctor_set_uint32(x_430, sizeof(void*)*4 + 4, x_426); +x_425 = lean_ctor_get_uint32(x_1, sizeof(void*)*5); +x_426 = lean_ctor_get(x_1, 1); +lean_inc(x_426); +x_427 = lean_ctor_get(x_1, 2); +lean_inc(x_427); +x_428 = lean_ctor_get(x_1, 3); +lean_inc(x_428); +x_429 = lean_ctor_get_uint32(x_1, sizeof(void*)*5 + 4); +x_430 = lean_ctor_get(x_1, 4); +lean_inc(x_430); +x_431 = l_Lean_LocalDecl_fvarId(x_418); +x_432 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_432, 0, x_431); +x_433 = l_Lean_Meta_SimpTheoremsArray_eraseTheorem(x_426, x_432); +x_434 = lean_alloc_ctor(0, 5, 8); +lean_ctor_set(x_434, 0, x_424); +lean_ctor_set(x_434, 1, x_433); +lean_ctor_set(x_434, 2, x_427); +lean_ctor_set(x_434, 3, x_428); +lean_ctor_set(x_434, 4, x_430); +lean_ctor_set_uint32(x_434, sizeof(void*)*5, x_425); +lean_ctor_set_uint32(x_434, sizeof(void*)*5 + 4, x_429); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_3); lean_inc(x_2); -lean_inc(x_419); -x_431 = l_Lean_Meta_simp(x_419, x_430, x_2, x_3, x_412, x_9, x_10, x_11, x_12, x_420); -if (lean_obj_tag(x_431) == 0) -{ -lean_object* x_432; lean_object* x_433; lean_object* x_434; -x_432 = lean_ctor_get(x_431, 0); -lean_inc(x_432); -x_433 = lean_ctor_get(x_432, 0); -lean_inc(x_433); -x_434 = lean_ctor_get(x_433, 1); -lean_inc(x_434); -if (lean_obj_tag(x_434) == 0) -{ -lean_object* x_435; lean_object* x_436; lean_object* x_437; lean_object* x_438; uint8_t x_439; -lean_dec(x_419); -lean_dec(x_415); -x_435 = lean_ctor_get(x_431, 1); -lean_inc(x_435); -lean_dec(x_431); -x_436 = lean_ctor_get(x_432, 1); +lean_inc(x_422); +x_435 = l_Lean_Meta_simp(x_422, x_434, x_2, x_3, x_414, x_9, x_10, x_11, x_12, x_423); +if (lean_obj_tag(x_435) == 0) +{ +lean_object* x_436; lean_object* x_437; lean_object* x_438; +x_436 = lean_ctor_get(x_435, 0); lean_inc(x_436); -if (lean_is_exclusive(x_432)) { - lean_ctor_release(x_432, 0); - lean_ctor_release(x_432, 1); - x_437 = x_432; -} else { - lean_dec_ref(x_432); - x_437 = lean_box(0); -} -x_438 = lean_ctor_get(x_433, 0); +x_437 = lean_ctor_get(x_436, 0); +lean_inc(x_437); +x_438 = lean_ctor_get(x_437, 1); lean_inc(x_438); -x_439 = l_Lean_Expr_isFalse(x_438); -if (x_439 == 0) +if (lean_obj_tag(x_438) == 0) { -lean_object* x_440; lean_object* x_441; -lean_dec(x_437); +lean_object* x_439; lean_object* x_440; lean_object* x_441; lean_object* x_442; uint8_t x_443; +lean_dec(x_422); +lean_dec(x_418); +x_439 = lean_ctor_get(x_435, 1); +lean_inc(x_439); +lean_dec(x_435); +x_440 = lean_ctor_get(x_436, 1); +lean_inc(x_440); +if (lean_is_exclusive(x_436)) { + lean_ctor_release(x_436, 0); + lean_ctor_release(x_436, 1); + x_441 = x_436; +} else { + lean_dec_ref(x_436); + x_441 = lean_box(0); +} +x_442 = lean_ctor_get(x_437, 0); +lean_inc(x_442); +x_443 = l_Lean_Expr_isFalse(x_442); +if (x_443 == 0) +{ +lean_object* x_444; lean_object* x_445; +lean_dec(x_441); +lean_dec(x_416); lean_dec(x_413); -lean_dec(x_410); -x_440 = lean_box(0); +x_444 = lean_box(0); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_4); -x_441 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_433, x_16, x_411, x_436, x_4, x_408, x_409, x_440, x_9, x_10, x_11, x_12, x_435); -if (lean_obj_tag(x_441) == 0) -{ -lean_object* x_442; lean_object* x_443; -x_442 = lean_ctor_get(x_441, 0); -lean_inc(x_442); -x_443 = lean_ctor_get(x_441, 1); -lean_inc(x_443); -lean_dec(x_441); -x_17 = x_442; -x_18 = x_443; +x_445 = l_Array_forInUnsafe_loop___at_Lean_Meta_simpGoal___spec__1___lambda__1(x_437, x_16, x_440, x_415, x_4, x_411, x_412, x_444, x_9, x_10, x_11, x_12, x_439); +if (lean_obj_tag(x_445) == 0) +{ +lean_object* x_446; lean_object* x_447; +x_446 = lean_ctor_get(x_445, 0); +lean_inc(x_446); +x_447 = lean_ctor_get(x_445, 1); +lean_inc(x_447); +lean_dec(x_445); +x_17 = x_446; +x_18 = x_447; goto block_25; } else { -lean_object* x_444; lean_object* x_445; lean_object* x_446; lean_object* x_447; +lean_object* x_448; lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -39592,108 +39633,108 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_444 = lean_ctor_get(x_441, 0); -lean_inc(x_444); -x_445 = lean_ctor_get(x_441, 1); -lean_inc(x_445); -if (lean_is_exclusive(x_441)) { - lean_ctor_release(x_441, 0); - lean_ctor_release(x_441, 1); - x_446 = x_441; +x_448 = lean_ctor_get(x_445, 0); +lean_inc(x_448); +x_449 = lean_ctor_get(x_445, 1); +lean_inc(x_449); +if (lean_is_exclusive(x_445)) { + lean_ctor_release(x_445, 0); + lean_ctor_release(x_445, 1); + x_450 = x_445; } else { - lean_dec_ref(x_441); - x_446 = lean_box(0); + lean_dec_ref(x_445); + x_450 = lean_box(0); } -if (lean_is_scalar(x_446)) { - x_447 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_450)) { + x_451 = lean_alloc_ctor(1, 2, 0); } else { - x_447 = x_446; + x_451 = x_450; } -lean_ctor_set(x_447, 0, x_444); -lean_ctor_set(x_447, 1, x_445); -return x_447; +lean_ctor_set(x_451, 0, x_448); +lean_ctor_set(x_451, 1, x_449); +return x_451; } } else { -lean_object* x_448; -lean_dec(x_433); -lean_inc(x_408); -x_448 = l_Lean_MVarId_getType(x_408, x_9, x_10, x_11, x_12, x_435); -if (lean_obj_tag(x_448) == 0) -{ -lean_object* x_449; lean_object* x_450; lean_object* x_451; lean_object* x_452; -x_449 = lean_ctor_get(x_448, 0); -lean_inc(x_449); -x_450 = lean_ctor_get(x_448, 1); -lean_inc(x_450); -lean_dec(x_448); -x_451 = l_Lean_Expr_fvar___override(x_16); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_452 = l_Lean_Meta_mkFalseElim(x_449, x_451, x_9, x_10, x_11, x_12, x_450); +lean_object* x_452; +lean_dec(x_437); +lean_inc(x_411); +x_452 = l_Lean_MVarId_getType(x_411, x_9, x_10, x_11, x_12, x_439); if (lean_obj_tag(x_452) == 0) { -lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; +lean_object* x_453; lean_object* x_454; lean_object* x_455; lean_object* x_456; x_453 = lean_ctor_get(x_452, 0); lean_inc(x_453); x_454 = lean_ctor_get(x_452, 1); lean_inc(x_454); lean_dec(x_452); -lean_inc(x_408); -x_455 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_408, x_453, x_9, x_10, x_11, x_12, x_454); -x_456 = lean_ctor_get(x_455, 1); -lean_inc(x_456); -lean_dec(x_455); -x_457 = lean_box(0); -lean_inc(x_436); -if (lean_is_scalar(x_437)) { - x_458 = lean_alloc_ctor(0, 2, 0); -} else { - x_458 = x_437; -} -lean_ctor_set(x_458, 0, x_457); -lean_ctor_set(x_458, 1, x_436); -x_459 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_459, 0, x_458); +x_455 = l_Lean_Expr_fvar___override(x_16); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +x_456 = l_Lean_Meta_mkFalseElim(x_453, x_455, x_9, x_10, x_11, x_12, x_454); +if (lean_obj_tag(x_456) == 0) +{ +lean_object* x_457; lean_object* x_458; lean_object* x_459; lean_object* x_460; lean_object* x_461; lean_object* x_462; lean_object* x_463; lean_object* x_464; lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; +x_457 = lean_ctor_get(x_456, 0); +lean_inc(x_457); +x_458 = lean_ctor_get(x_456, 1); +lean_inc(x_458); +lean_dec(x_456); +lean_inc(x_411); +x_459 = l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(x_411, x_457, x_9, x_10, x_11, x_12, x_458); +x_460 = lean_ctor_get(x_459, 1); +lean_inc(x_460); +lean_dec(x_459); +x_461 = lean_box(0); +lean_inc(x_440); +if (lean_is_scalar(x_441)) { + x_462 = lean_alloc_ctor(0, 2, 0); +} else { + x_462 = x_441; +} +lean_ctor_set(x_462, 0, x_461); +lean_ctor_set(x_462, 1, x_440); +x_463 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_463, 0, x_462); +if (lean_is_scalar(x_416)) { + x_464 = lean_alloc_ctor(0, 2, 0); +} else { + x_464 = x_416; +} +lean_ctor_set(x_464, 0, x_440); +lean_ctor_set(x_464, 1, x_415); if (lean_is_scalar(x_413)) { - x_460 = lean_alloc_ctor(0, 2, 0); -} else { - x_460 = x_413; -} -lean_ctor_set(x_460, 0, x_411); -lean_ctor_set(x_460, 1, x_436); -if (lean_is_scalar(x_410)) { - x_461 = lean_alloc_ctor(0, 2, 0); -} else { - x_461 = x_410; -} -lean_ctor_set(x_461, 0, x_409); -lean_ctor_set(x_461, 1, x_460); -x_462 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_462, 0, x_408); -lean_ctor_set(x_462, 1, x_461); -x_463 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_463, 0, x_459); -lean_ctor_set(x_463, 1, x_462); -x_464 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_464, 0, x_463); -x_17 = x_464; -x_18 = x_456; + x_465 = lean_alloc_ctor(0, 2, 0); +} else { + x_465 = x_413; +} +lean_ctor_set(x_465, 0, x_412); +lean_ctor_set(x_465, 1, x_464); +x_466 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_466, 0, x_411); +lean_ctor_set(x_466, 1, x_465); +x_467 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_467, 0, x_463); +lean_ctor_set(x_467, 1, x_466); +x_468 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_468, 0, x_467); +x_17 = x_468; +x_18 = x_460; goto block_25; } else { -lean_object* x_465; lean_object* x_466; lean_object* x_467; lean_object* x_468; -lean_dec(x_437); -lean_dec(x_436); +lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; +lean_dec(x_441); +lean_dec(x_440); +lean_dec(x_416); +lean_dec(x_415); lean_dec(x_413); +lean_dec(x_412); lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -39702,38 +39743,38 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_465 = lean_ctor_get(x_452, 0); -lean_inc(x_465); -x_466 = lean_ctor_get(x_452, 1); -lean_inc(x_466); -if (lean_is_exclusive(x_452)) { - lean_ctor_release(x_452, 0); - lean_ctor_release(x_452, 1); - x_467 = x_452; +x_469 = lean_ctor_get(x_456, 0); +lean_inc(x_469); +x_470 = lean_ctor_get(x_456, 1); +lean_inc(x_470); +if (lean_is_exclusive(x_456)) { + lean_ctor_release(x_456, 0); + lean_ctor_release(x_456, 1); + x_471 = x_456; } else { - lean_dec_ref(x_452); - x_467 = lean_box(0); + lean_dec_ref(x_456); + x_471 = lean_box(0); } -if (lean_is_scalar(x_467)) { - x_468 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_471)) { + x_472 = lean_alloc_ctor(1, 2, 0); } else { - x_468 = x_467; + x_472 = x_471; } -lean_ctor_set(x_468, 0, x_465); -lean_ctor_set(x_468, 1, x_466); -return x_468; +lean_ctor_set(x_472, 0, x_469); +lean_ctor_set(x_472, 1, x_470); +return x_472; } } else { -lean_object* x_469; lean_object* x_470; lean_object* x_471; lean_object* x_472; -lean_dec(x_437); -lean_dec(x_436); +lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; +lean_dec(x_441); +lean_dec(x_440); +lean_dec(x_416); +lean_dec(x_415); lean_dec(x_413); +lean_dec(x_412); lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_16); lean_dec(x_12); lean_dec(x_11); @@ -39743,177 +39784,177 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_469 = lean_ctor_get(x_448, 0); -lean_inc(x_469); -x_470 = lean_ctor_get(x_448, 1); -lean_inc(x_470); -if (lean_is_exclusive(x_448)) { - lean_ctor_release(x_448, 0); - lean_ctor_release(x_448, 1); - x_471 = x_448; +x_473 = lean_ctor_get(x_452, 0); +lean_inc(x_473); +x_474 = lean_ctor_get(x_452, 1); +lean_inc(x_474); +if (lean_is_exclusive(x_452)) { + lean_ctor_release(x_452, 0); + lean_ctor_release(x_452, 1); + x_475 = x_452; } else { - lean_dec_ref(x_448); - x_471 = lean_box(0); + lean_dec_ref(x_452); + x_475 = lean_box(0); } -if (lean_is_scalar(x_471)) { - x_472 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_475)) { + x_476 = lean_alloc_ctor(1, 2, 0); } else { - x_472 = x_471; + x_476 = x_475; } -lean_ctor_set(x_472, 0, x_469); -lean_ctor_set(x_472, 1, x_470); -return x_472; +lean_ctor_set(x_476, 0, x_473); +lean_ctor_set(x_476, 1, x_474); +return x_476; } } } else { -lean_object* x_473; lean_object* x_474; lean_object* x_475; lean_object* x_476; lean_object* x_477; uint8_t x_478; lean_object* x_479; -if (lean_is_exclusive(x_434)) { - lean_ctor_release(x_434, 0); - x_473 = x_434; +lean_object* x_477; lean_object* x_478; lean_object* x_479; lean_object* x_480; lean_object* x_481; uint8_t x_482; lean_object* x_483; +if (lean_is_exclusive(x_438)) { + lean_ctor_release(x_438, 0); + x_477 = x_438; } else { - lean_dec_ref(x_434); - x_473 = lean_box(0); + lean_dec_ref(x_438); + x_477 = lean_box(0); } -x_474 = lean_ctor_get(x_431, 1); -lean_inc(x_474); -lean_dec(x_431); -x_475 = lean_ctor_get(x_432, 1); -lean_inc(x_475); -if (lean_is_exclusive(x_432)) { - lean_ctor_release(x_432, 0); - lean_ctor_release(x_432, 1); - x_476 = x_432; -} else { - lean_dec_ref(x_432); - x_476 = lean_box(0); -} -x_477 = l_Lean_Expr_fvar___override(x_16); -x_478 = 1; +x_478 = lean_ctor_get(x_435, 1); +lean_inc(x_478); +lean_dec(x_435); +x_479 = lean_ctor_get(x_436, 1); +lean_inc(x_479); +if (lean_is_exclusive(x_436)) { + lean_ctor_release(x_436, 0); + lean_ctor_release(x_436, 1); + x_480 = x_436; +} else { + lean_dec_ref(x_436); + x_480 = lean_box(0); +} +x_481 = l_Lean_Expr_fvar___override(x_16); +x_482 = 1; lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -lean_inc(x_408); -x_479 = l_Lean_Meta_applySimpResultToProp(x_408, x_477, x_419, x_433, x_478, x_9, x_10, x_11, x_12, x_474); -lean_dec(x_419); -if (lean_obj_tag(x_479) == 0) +lean_inc(x_411); +x_483 = l_Lean_Meta_applySimpResultToProp(x_411, x_481, x_422, x_437, x_482, x_9, x_10, x_11, x_12, x_478); +lean_dec(x_422); +if (lean_obj_tag(x_483) == 0) { -lean_object* x_480; -x_480 = lean_ctor_get(x_479, 0); -lean_inc(x_480); -if (lean_obj_tag(x_480) == 0) +lean_object* x_484; +x_484 = lean_ctor_get(x_483, 0); +lean_inc(x_484); +if (lean_obj_tag(x_484) == 0) { -lean_object* x_481; lean_object* x_482; lean_object* x_483; lean_object* x_484; lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; -lean_dec(x_415); -x_481 = lean_ctor_get(x_479, 1); -lean_inc(x_481); -lean_dec(x_479); -x_482 = lean_box(0); -lean_inc(x_475); -if (lean_is_scalar(x_476)) { - x_483 = lean_alloc_ctor(0, 2, 0); +lean_object* x_485; lean_object* x_486; lean_object* x_487; lean_object* x_488; lean_object* x_489; lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; +lean_dec(x_418); +x_485 = lean_ctor_get(x_483, 1); +lean_inc(x_485); +lean_dec(x_483); +x_486 = lean_box(0); +lean_inc(x_479); +if (lean_is_scalar(x_480)) { + x_487 = lean_alloc_ctor(0, 2, 0); +} else { + x_487 = x_480; +} +lean_ctor_set(x_487, 0, x_486); +lean_ctor_set(x_487, 1, x_479); +if (lean_is_scalar(x_477)) { + x_488 = lean_alloc_ctor(1, 1, 0); } else { - x_483 = x_476; + x_488 = x_477; } -lean_ctor_set(x_483, 0, x_482); -lean_ctor_set(x_483, 1, x_475); -if (lean_is_scalar(x_473)) { - x_484 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_488, 0, x_487); +if (lean_is_scalar(x_416)) { + x_489 = lean_alloc_ctor(0, 2, 0); } else { - x_484 = x_473; + x_489 = x_416; } -lean_ctor_set(x_484, 0, x_483); +lean_ctor_set(x_489, 0, x_479); +lean_ctor_set(x_489, 1, x_415); if (lean_is_scalar(x_413)) { - x_485 = lean_alloc_ctor(0, 2, 0); -} else { - x_485 = x_413; -} -lean_ctor_set(x_485, 0, x_411); -lean_ctor_set(x_485, 1, x_475); -if (lean_is_scalar(x_410)) { - x_486 = lean_alloc_ctor(0, 2, 0); -} else { - x_486 = x_410; -} -lean_ctor_set(x_486, 0, x_409); -lean_ctor_set(x_486, 1, x_485); -x_487 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_487, 0, x_408); -lean_ctor_set(x_487, 1, x_486); -x_488 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_488, 0, x_484); -lean_ctor_set(x_488, 1, x_487); -x_489 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_489, 0, x_488); -x_17 = x_489; -x_18 = x_481; + x_490 = lean_alloc_ctor(0, 2, 0); +} else { + x_490 = x_413; +} +lean_ctor_set(x_490, 0, x_412); +lean_ctor_set(x_490, 1, x_489); +x_491 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_491, 0, x_411); +lean_ctor_set(x_491, 1, x_490); +x_492 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_492, 0, x_488); +lean_ctor_set(x_492, 1, x_491); +x_493 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_493, 0, x_492); +x_17 = x_493; +x_18 = x_485; goto block_25; } else { -lean_object* x_490; lean_object* x_491; lean_object* x_492; lean_object* x_493; lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; -lean_dec(x_476); -lean_dec(x_473); -x_490 = lean_ctor_get(x_480, 0); -lean_inc(x_490); +lean_object* x_494; lean_object* x_495; lean_object* x_496; lean_object* x_497; lean_object* x_498; lean_object* x_499; lean_object* x_500; lean_object* x_501; lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; lean_dec(x_480); -x_491 = lean_ctor_get(x_479, 1); -lean_inc(x_491); -lean_dec(x_479); -x_492 = lean_ctor_get(x_490, 0); -lean_inc(x_492); -x_493 = lean_ctor_get(x_490, 1); -lean_inc(x_493); -lean_dec(x_490); -x_494 = l_Lean_LocalDecl_userName(x_415); -lean_dec(x_415); -x_495 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_495, 0, x_494); -lean_ctor_set(x_495, 1, x_493); -lean_ctor_set(x_495, 2, x_492); -x_496 = lean_array_push(x_411, x_495); +lean_dec(x_477); +x_494 = lean_ctor_get(x_484, 0); +lean_inc(x_494); +lean_dec(x_484); +x_495 = lean_ctor_get(x_483, 1); +lean_inc(x_495); +lean_dec(x_483); +x_496 = lean_ctor_get(x_494, 0); +lean_inc(x_496); +x_497 = lean_ctor_get(x_494, 1); +lean_inc(x_497); +lean_dec(x_494); +x_498 = l_Lean_LocalDecl_userName(x_418); +lean_dec(x_418); +x_499 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_499, 0, x_498); +lean_ctor_set(x_499, 1, x_497); +lean_ctor_set(x_499, 2, x_496); +x_500 = lean_array_push(x_415, x_499); +if (lean_is_scalar(x_416)) { + x_501 = lean_alloc_ctor(0, 2, 0); +} else { + x_501 = x_416; +} +lean_ctor_set(x_501, 0, x_479); +lean_ctor_set(x_501, 1, x_500); if (lean_is_scalar(x_413)) { - x_497 = lean_alloc_ctor(0, 2, 0); + x_502 = lean_alloc_ctor(0, 2, 0); } else { - x_497 = x_413; + x_502 = x_413; } -lean_ctor_set(x_497, 0, x_496); -lean_ctor_set(x_497, 1, x_475); -if (lean_is_scalar(x_410)) { - x_498 = lean_alloc_ctor(0, 2, 0); -} else { - x_498 = x_410; -} -lean_ctor_set(x_498, 0, x_409); -lean_ctor_set(x_498, 1, x_497); -x_499 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_499, 0, x_408); -lean_ctor_set(x_499, 1, x_498); +lean_ctor_set(x_502, 0, x_412); +lean_ctor_set(x_502, 1, x_501); +x_503 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_503, 0, x_411); +lean_ctor_set(x_503, 1, x_502); lean_inc(x_4); -x_500 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_500, 0, x_4); -lean_ctor_set(x_500, 1, x_499); -x_501 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_501, 0, x_500); -x_17 = x_501; -x_18 = x_491; +x_504 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_504, 0, x_4); +lean_ctor_set(x_504, 1, x_503); +x_505 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_505, 0, x_504); +x_17 = x_505; +x_18 = x_495; goto block_25; } } else { -lean_object* x_502; lean_object* x_503; lean_object* x_504; lean_object* x_505; -lean_dec(x_476); -lean_dec(x_475); -lean_dec(x_473); +lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; +lean_dec(x_480); +lean_dec(x_479); +lean_dec(x_477); +lean_dec(x_418); +lean_dec(x_416); lean_dec(x_415); lean_dec(x_413); +lean_dec(x_412); lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -39922,39 +39963,39 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_502 = lean_ctor_get(x_479, 0); -lean_inc(x_502); -x_503 = lean_ctor_get(x_479, 1); -lean_inc(x_503); -if (lean_is_exclusive(x_479)) { - lean_ctor_release(x_479, 0); - lean_ctor_release(x_479, 1); - x_504 = x_479; +x_506 = lean_ctor_get(x_483, 0); +lean_inc(x_506); +x_507 = lean_ctor_get(x_483, 1); +lean_inc(x_507); +if (lean_is_exclusive(x_483)) { + lean_ctor_release(x_483, 0); + lean_ctor_release(x_483, 1); + x_508 = x_483; } else { - lean_dec_ref(x_479); - x_504 = lean_box(0); + lean_dec_ref(x_483); + x_508 = lean_box(0); } -if (lean_is_scalar(x_504)) { - x_505 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_508)) { + x_509 = lean_alloc_ctor(1, 2, 0); } else { - x_505 = x_504; + x_509 = x_508; } -lean_ctor_set(x_505, 0, x_502); -lean_ctor_set(x_505, 1, x_503); -return x_505; +lean_ctor_set(x_509, 0, x_506); +lean_ctor_set(x_509, 1, x_507); +return x_509; } } } else { -lean_object* x_506; lean_object* x_507; lean_object* x_508; lean_object* x_509; -lean_dec(x_419); +lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; +lean_dec(x_422); +lean_dec(x_418); +lean_dec(x_416); lean_dec(x_415); lean_dec(x_413); +lean_dec(x_412); lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_16); lean_dec(x_12); lean_dec(x_11); @@ -39964,37 +40005,37 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_506 = lean_ctor_get(x_431, 0); -lean_inc(x_506); -x_507 = lean_ctor_get(x_431, 1); -lean_inc(x_507); -if (lean_is_exclusive(x_431)) { - lean_ctor_release(x_431, 0); - lean_ctor_release(x_431, 1); - x_508 = x_431; +x_510 = lean_ctor_get(x_435, 0); +lean_inc(x_510); +x_511 = lean_ctor_get(x_435, 1); +lean_inc(x_511); +if (lean_is_exclusive(x_435)) { + lean_ctor_release(x_435, 0); + lean_ctor_release(x_435, 1); + x_512 = x_435; } else { - lean_dec_ref(x_431); - x_508 = lean_box(0); + lean_dec_ref(x_435); + x_512 = lean_box(0); } -if (lean_is_scalar(x_508)) { - x_509 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_512)) { + x_513 = lean_alloc_ctor(1, 2, 0); } else { - x_509 = x_508; + x_513 = x_512; } -lean_ctor_set(x_509, 0, x_506); -lean_ctor_set(x_509, 1, x_507); -return x_509; +lean_ctor_set(x_513, 0, x_510); +lean_ctor_set(x_513, 1, x_511); +return x_513; } } else { -lean_object* x_510; lean_object* x_511; lean_object* x_512; lean_object* x_513; +lean_object* x_514; lean_object* x_515; lean_object* x_516; lean_object* x_517; +lean_dec(x_416); +lean_dec(x_415); +lean_dec(x_414); lean_dec(x_413); lean_dec(x_412); lean_dec(x_411); -lean_dec(x_410); -lean_dec(x_409); -lean_dec(x_408); lean_dec(x_16); lean_dec(x_12); lean_dec(x_11); @@ -40004,26 +40045,26 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_510 = lean_ctor_get(x_414, 0); -lean_inc(x_510); -x_511 = lean_ctor_get(x_414, 1); -lean_inc(x_511); -if (lean_is_exclusive(x_414)) { - lean_ctor_release(x_414, 0); - lean_ctor_release(x_414, 1); - x_512 = x_414; +x_514 = lean_ctor_get(x_417, 0); +lean_inc(x_514); +x_515 = lean_ctor_get(x_417, 1); +lean_inc(x_515); +if (lean_is_exclusive(x_417)) { + lean_ctor_release(x_417, 0); + lean_ctor_release(x_417, 1); + x_516 = x_417; } else { - lean_dec_ref(x_414); - x_512 = lean_box(0); + lean_dec_ref(x_417); + x_516 = lean_box(0); } -if (lean_is_scalar(x_512)) { - x_513 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_516)) { + x_517 = lean_alloc_ctor(1, 2, 0); } else { - x_513 = x_512; + x_517 = x_516; } -lean_ctor_set(x_513, 0, x_510); -lean_ctor_set(x_513, 1, x_511); -return x_513; +lean_ctor_set(x_517, 0, x_514); +lean_ctor_set(x_517, 1, x_515); +return x_517; } } block_25: @@ -40514,8 +40555,8 @@ lean_dec(x_14); x_16 = lean_box(0); x_17 = l_Lean_Meta_Simp_lambdaTelescopeDSimp___rarg___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_3); +lean_ctor_set(x_18, 0, x_3); +lean_ctor_set(x_18, 1, x_17); x_19 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -40569,7 +40610,7 @@ x_35 = lean_ctor_get(x_29, 1); lean_inc(x_35); lean_dec(x_29); x_36 = lean_box(0); -x_37 = l_Lean_Meta_simpGoal___lambda__3(x_34, x_5, x_1, x_22, x_17, x_33, x_4, x_24, x_23, x_8, x_6, x_7, x_16, x_32, x_35, x_36, x_9, x_10, x_11, x_12, x_31); +x_37 = l_Lean_Meta_simpGoal___lambda__3(x_35, x_5, x_1, x_22, x_17, x_33, x_4, x_24, x_23, x_8, x_6, x_7, x_16, x_32, x_34, x_36, x_9, x_10, x_11, x_12, x_31); return x_37; } else @@ -40687,7 +40728,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_simpGoal(lean_object* x_1, lean_object* x_2 _start: { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = l_Lean_Meta_Simp_main___closed__1; +x_13 = l_Lean_Meta_Simp_main_simpMain___closed__1; x_14 = lean_box(x_5); lean_inc(x_1); x_15 = lean_alloc_closure((void*)(l_Lean_Meta_simpGoal___lambda__4___boxed), 13, 8); @@ -40874,37 +40915,39 @@ lean_dec(x_14); x_17 = !lean_is_exclusive(x_4); 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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; x_18 = lean_ctor_get(x_4, 0); x_19 = lean_ctor_get(x_4, 1); x_20 = lean_ctor_get(x_4, 2); x_21 = lean_ctor_get(x_4, 3); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_12); +x_22 = lean_ctor_get(x_4, 4); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_12); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_23 = l_Lean_Meta_SimpTheoremsArray_addTheorem(x_19, x_22, x_16, x_5, x_6, x_7, x_8, x_15); -if (lean_obj_tag(x_23) == 0) +x_24 = l_Lean_Meta_SimpTheoremsArray_addTheorem(x_19, x_23, x_16, x_5, x_6, x_7, x_8, x_15); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; size_t x_26; size_t x_27; -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; size_t x_27; size_t x_28; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -lean_ctor_set(x_4, 1, x_24); -x_26 = 1; -x_27 = lean_usize_add(x_3, x_26); -x_3 = x_27; -x_9 = x_25; +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +lean_ctor_set(x_4, 1, x_25); +x_27 = 1; +x_28 = lean_usize_add(x_3, x_27); +x_3 = x_28; +x_9 = x_26; goto _start; } else { -uint8_t x_29; +uint8_t x_30; lean_free_object(x_4); +lean_dec(x_22); lean_dec(x_21); lean_dec(x_20); lean_dec(x_18); @@ -40912,128 +40955,132 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_29 = !lean_is_exclusive(x_23); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_24); +if (x_30 == 0) { -return x_23; +return x_24; } 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_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_24, 0); +x_32 = lean_ctor_get(x_24, 1); +lean_inc(x_32); 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; +lean_dec(x_24); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } else { -lean_object* x_33; uint32_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint32_t x_38; lean_object* x_39; lean_object* x_40; -x_33 = lean_ctor_get(x_4, 0); -x_34 = lean_ctor_get_uint32(x_4, sizeof(void*)*4); -x_35 = lean_ctor_get(x_4, 1); -x_36 = lean_ctor_get(x_4, 2); -x_37 = lean_ctor_get(x_4, 3); -x_38 = lean_ctor_get_uint32(x_4, sizeof(void*)*4 + 4); +lean_object* x_34; uint32_t x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; uint32_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_34 = lean_ctor_get(x_4, 0); +x_35 = lean_ctor_get_uint32(x_4, sizeof(void*)*5); +x_36 = lean_ctor_get(x_4, 1); +x_37 = lean_ctor_get(x_4, 2); +x_38 = lean_ctor_get(x_4, 3); +x_39 = lean_ctor_get_uint32(x_4, sizeof(void*)*5 + 4); +x_40 = lean_ctor_get(x_4, 4); +lean_inc(x_40); +lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); -lean_inc(x_35); -lean_inc(x_33); +lean_inc(x_34); lean_dec(x_4); -x_39 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_39, 0, x_12); +x_41 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_41, 0, x_12); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_40 = l_Lean_Meta_SimpTheoremsArray_addTheorem(x_35, x_39, x_16, x_5, x_6, x_7, x_8, x_15); -if (lean_obj_tag(x_40) == 0) +x_42 = l_Lean_Meta_SimpTheoremsArray_addTheorem(x_36, x_41, x_16, x_5, x_6, x_7, x_8, x_15); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; size_t x_44; size_t x_45; -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_alloc_ctor(0, 4, 8); -lean_ctor_set(x_43, 0, x_33); -lean_ctor_set(x_43, 1, x_41); -lean_ctor_set(x_43, 2, x_36); -lean_ctor_set(x_43, 3, x_37); -lean_ctor_set_uint32(x_43, sizeof(void*)*4, x_34); -lean_ctor_set_uint32(x_43, sizeof(void*)*4 + 4, x_38); -x_44 = 1; -x_45 = lean_usize_add(x_3, x_44); -x_3 = x_45; -x_4 = x_43; -x_9 = x_42; +lean_object* x_43; lean_object* x_44; lean_object* x_45; size_t x_46; size_t x_47; +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_alloc_ctor(0, 5, 8); +lean_ctor_set(x_45, 0, x_34); +lean_ctor_set(x_45, 1, x_43); +lean_ctor_set(x_45, 2, x_37); +lean_ctor_set(x_45, 3, x_38); +lean_ctor_set(x_45, 4, x_40); +lean_ctor_set_uint32(x_45, sizeof(void*)*5, x_35); +lean_ctor_set_uint32(x_45, sizeof(void*)*5 + 4, x_39); +x_46 = 1; +x_47 = lean_usize_add(x_3, x_46); +x_3 = x_47; +x_4 = x_45; +x_9 = x_44; goto _start; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_40); +lean_dec(x_38); lean_dec(x_37); -lean_dec(x_36); -lean_dec(x_33); +lean_dec(x_34); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_47 = lean_ctor_get(x_40, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_40, 1); -lean_inc(x_48); -if (lean_is_exclusive(x_40)) { - lean_ctor_release(x_40, 0); - lean_ctor_release(x_40, 1); - x_49 = x_40; +x_49 = lean_ctor_get(x_42, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_42, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_42)) { + lean_ctor_release(x_42, 0); + lean_ctor_release(x_42, 1); + x_51 = x_42; } else { - lean_dec_ref(x_40); - x_49 = lean_box(0); + lean_dec_ref(x_42); + x_51 = lean_box(0); } -if (lean_is_scalar(x_49)) { - x_50 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); } else { - x_50 = x_49; + x_52 = x_51; } -lean_ctor_set(x_50, 0, x_47); -lean_ctor_set(x_50, 1, x_48); -return x_50; +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } } else { -uint8_t x_51; +uint8_t x_53; lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_51 = !lean_is_exclusive(x_13); -if (x_51 == 0) +x_53 = !lean_is_exclusive(x_13); +if (x_53 == 0) { return x_13; } 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); -lean_inc(x_53); -lean_inc(x_52); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_13, 0); +x_55 = lean_ctor_get(x_13, 1); +lean_inc(x_55); +lean_inc(x_54); lean_dec(x_13); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_52); -lean_ctor_set(x_54, 1, x_53); -return x_54; +x_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; } } } @@ -41709,6 +41756,43 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_Meta_Simp_withNewLemmas___rarg___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4; +x_3 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_1); +lean_ctor_set(x_3, 2, x_2); +return x_3; +} +} +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3; +x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__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; +} +} LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { @@ -41760,7 +41844,7 @@ lean_inc(x_33); x_34 = lean_ctor_get(x_32, 1); lean_inc(x_34); lean_dec(x_32); -x_35 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3; +x_35 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -42239,7 +42323,7 @@ lean_inc(x_118); x_119 = lean_ctor_get(x_117, 1); lean_inc(x_119); lean_dec(x_117); -x_120 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3; +x_120 = l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); @@ -43506,7 +43590,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_dsimpGoal(lean_object* x_1, lean_object* x_ _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = l_Lean_Meta_Simp_main___closed__1; +x_12 = l_Lean_Meta_Simp_main_simpMain___closed__1; x_13 = lean_box(x_4); lean_inc(x_1); x_14 = lean_alloc_closure((void*)(l_Lean_Meta_dsimpGoal___lambda__7___boxed), 12, 7); @@ -43630,6 +43714,7 @@ lean_object* initialize_Lean_Meta_Transform(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Replace(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_UnifyEq(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Tactic_Simp_Rewrite(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Meta_Tactic_Simp_Diagnostics(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Match_Value(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_Simp_Main(uint8_t builtin, lean_object* w) { @@ -43648,6 +43733,9 @@ lean_dec_ref(res); res = initialize_Lean_Meta_Tactic_Simp_Rewrite(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Meta_Tactic_Simp_Diagnostics(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Match_Value(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -43906,10 +43994,8 @@ l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___closed__1 = _init_l_Lean_ lean_mark_persistent(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___closed__1); l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___closed__2 = _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___closed__2(); lean_mark_persistent(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__5___closed__2); -l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1 = _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__1); -l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2 = _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___closed__2); +l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___closed__1 = _init_l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_trySimpCongrTheorem_x3f___lambda__9___closed__1); l_Lean_Meta_Simp_congr___closed__1 = _init_l_Lean_Meta_Simp_congr___closed__1(); lean_mark_persistent(l_Lean_Meta_Simp_congr___closed__1); l_Lean_Meta_Simp_simpStep___closed__1 = _init_l_Lean_Meta_Simp_simpStep___closed__1(); @@ -43928,18 +44014,16 @@ l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__1 = _init_l_Lean_Meta_Simp_si lean_mark_persistent(l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__1); l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__2 = _init_l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__2(); lean_mark_persistent(l_Lean_Meta_Simp_simpImpl_go___lambda__2___closed__2); +l_Lean_Meta_Simp_main_simpMain___closed__1 = _init_l_Lean_Meta_Simp_main_simpMain___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_main_simpMain___closed__1); l_Lean_Meta_Simp_main___closed__1 = _init_l_Lean_Meta_Simp_main___closed__1(); lean_mark_persistent(l_Lean_Meta_Simp_main___closed__1); l_Lean_Meta_Simp_main___closed__2 = _init_l_Lean_Meta_Simp_main___closed__2(); lean_mark_persistent(l_Lean_Meta_Simp_main___closed__2); -l_Lean_Meta_Simp_main___closed__3 = _init_l_Lean_Meta_Simp_main___closed__3(); -lean_mark_persistent(l_Lean_Meta_Simp_main___closed__3); -l_Lean_Meta_Simp_dsimpMain___closed__1 = _init_l_Lean_Meta_Simp_dsimpMain___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_dsimpMain___closed__1); -l_Lean_Meta_Simp_dsimpMain___closed__2 = _init_l_Lean_Meta_Simp_dsimpMain___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_dsimpMain___closed__2); l_Lean_Meta_simp___closed__1 = _init_l_Lean_Meta_simp___closed__1(); lean_mark_persistent(l_Lean_Meta_simp___closed__1); +l_Lean_Meta_dsimp___closed__1 = _init_l_Lean_Meta_dsimp___closed__1(); +lean_mark_persistent(l_Lean_Meta_dsimp___closed__1); l_Lean_Meta_simpTargetCore___closed__1 = _init_l_Lean_Meta_simpTargetCore___closed__1(); lean_mark_persistent(l_Lean_Meta_simpTargetCore___closed__1); l_Lean_Meta_simpTargetCore___closed__2 = _init_l_Lean_Meta_simpTargetCore___closed__2(); @@ -43960,6 +44044,12 @@ l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__2 = _init_ lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__2); l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__3); +l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__5 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__5(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__5); +l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6 = _init_l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Meta_dsimpGoal___spec__1___closed__6); l_Lean_Meta_dsimpGoal___lambda__2___closed__1 = _init_l_Lean_Meta_dsimpGoal___lambda__2___closed__1(); lean_mark_persistent(l_Lean_Meta_dsimpGoal___lambda__2___closed__1); l_Lean_Meta_dsimpGoal___lambda__2___closed__2 = _init_l_Lean_Meta_dsimpGoal___lambda__2___closed__2(); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c index 14f9d93c9ec9..92dc70e119e9 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Rewrite.c @@ -14,11 +14,9 @@ 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_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*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewritePre___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_rewritePost(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___closed__4; @@ -39,11 +37,13 @@ extern lean_object* l_Lean_profiler; LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_Simp_simpMatch___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_Meta_Simp_discharge_x3f_x27___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f_go_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeDefault_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_tryTheoremWithExtraArgs_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore___closed__1; +static double l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Simp_discharge_x3f_x27___spec__5(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___spec__1___closed__4; static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__1___closed__1; @@ -57,7 +57,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simpMatchC static lean_object* l_Lean_Meta_Simp_mkDefaultMethodsCore___closed__1; lean_object* l_Lean_Meta_mkNoConfusion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f___closed__3; -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_simpGround___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_Meta_Simp_simpMatchCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_SimpTheoremsArray_isErased(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_DischargeResult_ofNat___boxed(lean_object*); @@ -73,15 +72,14 @@ static lean_object* l_panic___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean uint32_t l_UInt32_ofNatTruncate(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_synthesizeArgs___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getDecl(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeUsingAssumption_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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* l_Lean_ConstantInfo_levelParams(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkDefaultMethods___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__18; -static lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Meta_Simp_discharge_x3f_x27___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_Meta_Simp_simpMatch(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_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___closed__4; @@ -89,13 +87,12 @@ static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___closed__5; 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_Meta_Simp_isEqnThmHypothesis_go___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpUsingDecide___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_LocalContext_findDeclRevM_x3f___at_Lean_Meta_Simp_dischargeUsingAssumption_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_Lean_Meta_Simp_simpGround___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Meta_Linear_isLinearCnstr(lean_object*); lean_object* lean_array_fswap(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isEqnThmHypothesis___boxed(lean_object*); -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_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_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Simp_discharge_x3f_x27___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpGround___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__4___closed__1; @@ -139,11 +136,11 @@ uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_DischargeResult_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpMatchDiscrs_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___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_MonadExcept_ofExcept___at_Lean_Meta_Simp_simpGround___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_Meta_Simp_sevalGround___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* l_Lean_Expr_stripArgsN(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_insertionSort_traverse___at_Lean_Meta_Simp_rewrite_x3f___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_simpMatchDiscrs_x3f___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_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_preSEval___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_Expr_cleanupAnnotations(lean_object*); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__2; @@ -151,8 +148,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeDefault_x3f___lambda__2(lean_ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppArgsN_loop(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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_EXPORT lean_object* l_Lean_Meta_Simp_sevalGround___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_preSEval___lambda__1___closed__1; lean_object* l_Lean_Meta_Simp_getSimprocs___rarg(lean_object*, lean_object*); @@ -164,25 +161,23 @@ static lean_object* l_Lean_Meta_Simp_dpreDefault___closed__1; lean_object* l_Lean_Meta_getSimpCongrTheorems___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpArith___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_synthesizeArgs___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__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*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore___closed__3; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_instDecidableEqDischargeResult(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___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_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_rewritePre(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_PersistentArray_findSomeRevM_x3f___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___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_Meta_Simp_preDefault___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_Simp_discharge_x3f_x27___spec__7___rarg___closed__3; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpMatchCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_LocalDecl_index(lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); static lean_object* l_Lean_Meta_Simp_seval___closed__8; static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__1___closed__4; lean_object* l_Lean_Meta_mkEqFalse_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8___closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_Simp_discharge_x3f_x27___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Simp_simprocs; LEAN_EXPORT lean_object* l_Lean_hasAssignableMVar___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -197,15 +192,14 @@ LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_drewritePre___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*); size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_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_Meta_Simp_dischargeDefault_x3f___lambda__1(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_Meta_Simp_isEqnThmHypothesis_go(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpGround___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_sevalGround(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_Simp_discharge_x3f_x27___spec__6___lambda__5(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*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___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_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_simpGround___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_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___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* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_Simp_simpCtorEq___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*); @@ -245,11 +239,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_rewritePost___boxed(lean_object*, lean lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_isLevelMVarAssignable___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_rewrite_x3f___closed__3; static lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___closed__7; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkMethods___elambda__3(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_isEq(lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___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_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_seval___closed__2; static lean_object* l_Lean_Meta_Simp_seval___closed__4; static lean_object* l_Lean_Meta_Simp_sevalGround___lambda__2___closed__2; @@ -259,7 +254,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_Simp_tryThe LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_Simp_discharge_x3f_x27___spec__7___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_Simp_discharge_x3f_x27___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___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___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore___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_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___closed__3; lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); uint8_t l_Lean_Expr_isMVar(lean_object*); static lean_object* l_Lean_Meta_Simp_simpCtorEq___lambda__1___closed__3; @@ -276,12 +273,11 @@ lean_object* l_Lean_Meta_getEqnsFor_x3f(lean_object*, uint8_t, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpMatch___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*); extern lean_object* l_Lean_checkEmoji; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_postDefault___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_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewrite_x3f___spec__3___closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_Simp_tryTheoremWithExtraArgs_x3f___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_Simp_discharge_x3f_x27___spec__6___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_postDefault___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeDefault_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -291,20 +287,23 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_DischargeResult_noConfusion___rarg(uin static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewritePre___spec__1___closed__1; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_DischargeResult_ofNat(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore___closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, 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_Simp_seval___closed__7; static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Simp_discharge_x3f_x27___spec__3___closed__2; static lean_object* l_Lean_isLevelMVarAssignable___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__5___closed__3; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___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_Meta_Simp_DischargeResult_noConfusion___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__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*); extern lean_object* l_Lean_levelZero; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f___closed__5; lean_object* lean_io_mono_nanos_now(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_rewrite_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_LocalContext_findDeclRevM_x3f___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Level_hasMVar(lean_object*); static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__1___closed__8; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___rarg(lean_object*, lean_object*); lean_object* l_Lean_Meta_whnfD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -316,6 +315,7 @@ static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Si static lean_object* l_Lean_Meta_Simp_simpCtorEq___closed__5; static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__3___closed__5; uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_simpGround___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkDecide(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_drewritePost___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*); extern lean_object* l_Lean_Meta_tactic_skipAssignedInstances; @@ -328,7 +328,7 @@ static lean_object* l_Lean_isLevelMVarAssignable___at___private_Lean_Meta_Tactic LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpArith___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_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_userPostSimprocs(lean_object*, 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_Simp_discharge_x3f_x27___spec__6___lambda__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_SimprocEntry_try___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* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___closed__8; @@ -341,6 +341,7 @@ uint8_t l_Lean_Meta_Linear_parentIsTarget(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSEvalContext(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_getMatchWithExtra___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_dsimp(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_threshold; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___closed__6; static lean_object* l_Lean_Meta_Simp_postDefault___lambda__3___closed__1; @@ -358,18 +359,19 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkMethods___elambda__2(lean_object*, l extern lean_object* l_Lean_Meta_instMonadMetaM; static lean_object* l_Lean_Meta_Simp_simpCtorEq___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpGround___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_PersistentArray_findSomeRevMAux___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MVarId_isAssignable___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__10; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simpMatchCore___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_postDefault___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_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_Simp_simpGround___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__7___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___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*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpCtorEq___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_postSEval(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_instBEqOrigin___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_sevalGround___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_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___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_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpGround___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_Std_Range_forIn_loop___at_Lean_Meta_Simp_simpMatchDiscrs_x3f___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instDecidableEqDischargeResult___boxed(lean_object*, lean_object*); @@ -383,21 +385,23 @@ lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTr 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*); +static double l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__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*, lean_object*); lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_drewritePost___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Meta_Simp_tryTheoremWithExtraArgs_x3f___spec__1___rarg___boxed(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*); extern lean_object* l_Lean_crossEmoji; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_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_Lean_Meta_Simp_recordTriedSimpTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppRevArgsAux(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_postDefault___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_LocalContext_findDeclRevM_x3f___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isFalse(lean_object*); static lean_object* l_Lean_Meta_ppOrigin___at_Lean_Meta_Simp_discharge_x3f_x27___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpMatchDiscrs_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*); @@ -408,19 +412,17 @@ static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__15; lean_object* l_Lean_Meta_mkOfEqTrue(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_Lean_Meta_Simp_mkSEvalContext___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_simpMatchDiscrs_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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_EXPORT lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_synthesizeArgs___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_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___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_Meta_Simp_postDefault___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevM_x3f___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___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_Meta_Simp_simpUsingDecide___lambda__1___closed__6; static lean_object* l_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___closed__3; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Meta_ACLt_main_lt(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_recordSimpTheorem(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__4___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_preDefault___lambda__1___closed__1; static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__11; static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__3; @@ -461,7 +463,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_ 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_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; lean_object* l_Lean_LocalDecl_type(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -482,7 +483,7 @@ static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Simp_discharge_x3f_x27___ 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 double l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__1; +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_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*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_mkSEvalContext___closed__2; @@ -496,21 +497,20 @@ lean_object* l_Lean_Meta_getFunInfoNArgs(lean_object*, lean_object*, lean_object static lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__7; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkMethods(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkMethods(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeGround(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_synthesizeArgs_synthesizeInstance___closed__7; 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*); uint8_t l_Lean_Expr_isTrue(lean_object*); lean_object* l_Lean_instantiateMVarsCore(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f_go_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_Simp_simpMatch___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_Meta_Simp_simpUsingDecide___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_Simp_preDefault(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_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__6___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_discharge_x3f_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_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___lambda__5(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*, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpCtorEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instHashableOrigin___boxed(lean_object*); @@ -537,10 +537,10 @@ static lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27__ LEAN_EXPORT 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_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSEvalMethods___rarg___boxed(lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___closed__5; 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); @@ -548,7 +548,10 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_ static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Lean_isLevelMVarAssignable___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpCtorEq___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___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_Meta_Simp_mkMethods___elambda__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_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__2; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkMethods___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_synthesizeArgs_synthesizeInstance___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkDefaultMethods(lean_object*, lean_object*, lean_object*); @@ -561,6 +564,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___s static lean_object* l_Lean_getConstInfo___at_Lean_Meta_Simp_discharge_x3f_x27___spec__3___closed__3; lean_object* lean_array_get_size(lean_object*); lean_object* l_Lean_LocalDecl_toExpr(lean_object*); +LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__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*, lean_object*, lean_object*); static lean_object* l_Lean_isLevelMVarAssignable___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__5___closed__2; LEAN_EXPORT lean_object* l_List_anyM___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___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___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -578,19 +582,21 @@ static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___closed__1; lean_object* lean_nat_add(lean_object*, lean_object*); uint8_t l_Lean_Expr_isConst(lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_Simp_simpGround___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_dischargeEqnThmHypothesis_x3f___closed__1; static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__16; static lean_object* l_Lean_Meta_Simp_mkDefaultMethods___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewritePost___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___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*); uint8_t l_Lean_Expr_isForall(lean_object*); static lean_object* l_Lean_Meta_Simp_dischargeEqnThmHypothesis_x3f___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dpreDefault(lean_object*, 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_Simp_simpGround___spec__1___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___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_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpArith___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_LocalContext_findDeclRevM_x3f___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_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_Array_findSomeRevM_x3f_find___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_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*); @@ -600,7 +606,6 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGroun lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpMatch___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___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___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_withTraceNode___at_Lean_Meta_Simp_simpGround___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_simp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -612,7 +617,6 @@ LEAN_EXPORT lean_object* l_Lean_hasAssignableLevelMVar___at___private_Lean_Meta_ lean_object* l_ReaderT_instMonad___rarg(lean_object*); static lean_object* l_Lean_Meta_Simp_discharge_x3f_x27___lambda__1___closed__7; static lean_object* l_Lean_Meta_ppSimpTheorem___at___private_Lean_Meta_Tactic_Simp_Rewrite_0__Lean_Meta_Simp_tryTheoremCore_go___spec__1___closed__6; -LEAN_EXPORT lean_object* l_Lean_PersistentArray_findSomeRevMAux___at_Lean_Meta_Simp_dischargeUsingAssumption_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewritePost___spec__1___closed__1; static lean_object* l_Lean_Meta_Simp_postSEval___lambda__1___closed__3; @@ -623,6 +627,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_sevalGround___s uint8_t l_Lean_Expr_hasExprMVar(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_rewrite_x3f___spec__3___closed__5; static lean_object* l_Lean_Meta_Simp_simpUsingDecide___lambda__1___closed__8; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_simpGround___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*); uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* l_Lean_Meta_Simp_simpCtorEq___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_DischargeResult_toCtorIdx(uint8_t x_1) { @@ -1536,119 +1541,7 @@ x_7 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTra return x_7; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___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) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_io_mono_nanos_now(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_apply_8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, 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; -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_io_mono_nanos_now(x_15); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -lean_object* x_18; uint8_t x_19; lean_object* x_20; double x_21; double x_22; double x_23; double x_24; double x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_18 = lean_ctor_get(x_16, 0); -x_19 = 0; -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Float_ofScientific(x_11, x_19, x_20); -lean_dec(x_11); -x_22 = l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8___closed__1; -x_23 = lean_float_div(x_21, x_22); -x_24 = l_Float_ofScientific(x_18, x_19, x_20); -lean_dec(x_18); -x_25 = lean_float_div(x_24, x_22); -x_26 = lean_box_float(x_23); -x_27 = lean_box_float(x_25); -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_14); -lean_ctor_set(x_29, 1, x_28); -lean_ctor_set(x_16, 0, x_29); -return x_16; -} -else -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; double x_34; double x_35; double x_36; double x_37; double x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_30 = lean_ctor_get(x_16, 0); -x_31 = lean_ctor_get(x_16, 1); -lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_16); -x_32 = 0; -x_33 = lean_unsigned_to_nat(0u); -x_34 = l_Float_ofScientific(x_11, x_32, x_33); -lean_dec(x_11); -x_35 = l_Lean_withStartStopSeconds___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8___closed__1; -x_36 = lean_float_div(x_34, x_35); -x_37 = l_Float_ofScientific(x_30, x_32, x_33); -lean_dec(x_30); -x_38 = lean_float_div(x_37, x_35); -x_39 = lean_box_float(x_36); -x_40 = lean_box_float(x_38); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_14); -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_31); -return x_43; -} -} -else -{ -uint8_t x_44; -lean_dec(x_11); -x_44 = !lean_is_exclusive(x_13); -if (x_44 == 0) -{ -return x_13; -} -else -{ -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___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, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; @@ -1917,7 +1810,7 @@ return x_102; } } } -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) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___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, 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) @@ -1942,82 +1835,16 @@ return x_13; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -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) -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_10, 0); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_12); -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); -x_16 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_16, 0, x_14); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -return x_17; -} -} -else -{ -uint8_t x_18; -x_18 = !lean_is_exclusive(x_10); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_10, 0); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set_tag(x_10, 0); -lean_ctor_set(x_10, 0, x_20); -return x_10; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -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(0, 1, 0); -lean_ctor_set(x_23, 0, x_21); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_inc(x_12); -x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__9(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +x_15 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__8(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); x_16 = lean_ctor_get(x_15, 1); lean_inc(x_16); lean_dec(x_15); -x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__10(x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_16); +x_17 = l_MonadExcept_ofExcept___at_Lean_Meta_Simp_discharge_x3f_x27___spec__9(x_4, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_16); lean_dec(x_12); lean_dec(x_9); lean_dec(x_7); @@ -2025,7 +1852,7 @@ lean_dec(x_4); return x_17; } } -static double _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__1() { +static double _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -2035,7 +1862,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__2() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -2043,11 +1870,11 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17, lean_object* x_18, lean_object* x_19) { _start: { double x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; -x_20 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__1; +x_20 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_21 = lean_alloc_ctor(0, 2, 17); @@ -2056,7 +1883,7 @@ lean_ctor_set(x_21, 1, x_3); lean_ctor_set_float(x_21, sizeof(void*)*2, x_20); lean_ctor_set_float(x_21, sizeof(void*)*2 + 8, x_20); lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 16, x_2); -x_22 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__2; +x_22 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2___closed__2; x_23 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_22); lean_dec(x_7); if (x_23 == 0) @@ -2067,7 +1894,7 @@ lean_object* x_24; lean_object* x_25; lean_dec(x_3); lean_dec(x_1); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(x_4, x_5, x_11, x_6, x_21, x_24, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +x_25 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_21, x_24, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); return x_25; } else @@ -2081,7 +1908,7 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +x_28 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); return x_28; } } @@ -2096,12 +1923,12 @@ lean_ctor_set_float(x_29, sizeof(void*)*2, x_9); lean_ctor_set_float(x_29, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_29, sizeof(void*)*2 + 16, x_2); x_30 = lean_box(0); -x_31 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__2(x_4, x_5, x_11, x_6, x_29, x_30, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +x_31 = l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_29, x_30, x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_19); return x_31; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_Meta_Simp_discharge_x3f_x27___spec__6___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -2109,16 +1936,16 @@ x_1 = lean_mk_string_from_bytes("", 15); return x_1; } } -static lean_object* _init_l_Lean_PPContext_runCoreM___rarg___closed__19() { +LEAN_EXPORT lean_object* l_Lean_PPContext_runCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("internal exception #", 20); -return x_1; +lean_object* x_4; lean_object* x_17; lean_object* x_18; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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; uint8_t x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; uint8_t x_71; +x_39 = lean_ctor_get(x_1, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_1, 3); +lean_inc(x_40); +x_41 = lean_ctor_get(x_1, 4); +lean_inc(x_41); +x_42 = lean_ctor_get(x_1, 5); +lean_inc(x_42); +lean_dec(x_1); +x_43 = l_Lean_Core_getMaxHeartbeats(x_40); +x_44 = l_Lean_PPContext_runCoreM___rarg___closed__5; +x_45 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_40, x_44); +x_46 = l_Lean_PPContext_runCoreM___rarg___closed__6; +x_47 = l_Lean_PPContext_runCoreM___rarg___closed__9; +x_48 = l_Lean_PPContext_runCoreM___rarg___closed__12; +x_49 = l_Lean_PPContext_runCoreM___rarg___closed__16; +x_50 = l_Lean_PPContext_runCoreM___rarg___closed__18; +x_51 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_51, 0, x_39); +lean_ctor_set(x_51, 1, x_46); +lean_ctor_set(x_51, 2, x_47); +lean_ctor_set(x_51, 3, x_48); +lean_ctor_set(x_51, 4, x_49); +lean_ctor_set(x_51, 5, x_48); +lean_ctor_set(x_51, 6, x_50); +x_52 = lean_io_get_num_heartbeats(x_3); +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_PPContext_runCoreM___rarg___closed__19; +x_56 = l_Lean_PPContext_runCoreM___rarg___closed__4; +x_57 = lean_unsigned_to_nat(0u); +x_58 = lean_unsigned_to_nat(1000u); +x_59 = lean_box(0); +x_60 = l_Lean_firstFrontendMacroScope; +x_61 = 0; +lean_inc(x_40); +x_62 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_62, 0, x_55); +lean_ctor_set(x_62, 1, x_56); +lean_ctor_set(x_62, 2, x_40); +lean_ctor_set(x_62, 3, x_57); +lean_ctor_set(x_62, 4, x_58); +lean_ctor_set(x_62, 5, x_59); +lean_ctor_set(x_62, 6, x_41); +lean_ctor_set(x_62, 7, x_42); +lean_ctor_set(x_62, 8, x_53); +lean_ctor_set(x_62, 9, x_43); +lean_ctor_set(x_62, 10, x_60); +lean_ctor_set_uint8(x_62, sizeof(void*)*11, x_61); +lean_ctor_set_uint8(x_62, sizeof(void*)*11 + 1, x_45); +x_63 = lean_st_mk_ref(x_51, x_54); +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_63, 1); +lean_inc(x_65); +lean_dec(x_63); +x_66 = lean_st_ref_get(x_64, x_65); +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, 0); +lean_inc(x_69); +lean_dec(x_67); +x_70 = l_Lean_Kernel_isDiagnosticsEnabled(x_69); +lean_dec(x_69); +if (x_70 == 0) +{ +if (x_45 == 0) +{ +uint8_t x_132; +x_132 = 1; +x_71 = x_132; +goto block_131; } +else +{ +x_71 = x_61; +goto block_131; } -LEAN_EXPORT lean_object* l_Lean_PPContext_runCoreM___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: +} +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; 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; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_4 = lean_ctor_get(x_1, 0); -x_5 = lean_ctor_get(x_1, 3); -x_6 = lean_ctor_get(x_1, 4); -x_7 = lean_ctor_get(x_1, 5); -x_8 = l_Lean_Core_getMaxHeartbeats(x_5); -x_9 = l_Lean_PPContext_runCoreM___rarg___closed__4; -x_10 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_9); -x_11 = l_Lean_PPContext_runCoreM___rarg___closed__5; -x_12 = l_Lean_PPContext_runCoreM___rarg___closed__8; -x_13 = l_Lean_PPContext_runCoreM___rarg___closed__11; -x_14 = l_Lean_PPContext_runCoreM___rarg___closed__15; -x_15 = l_Lean_PPContext_runCoreM___rarg___closed__17; -lean_inc(x_4); -x_16 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_16, 0, x_4); -lean_ctor_set(x_16, 1, x_11); -lean_ctor_set(x_16, 2, x_12); -lean_ctor_set(x_16, 3, x_13); -lean_ctor_set(x_16, 4, x_14); -lean_ctor_set(x_16, 5, x_13); -lean_ctor_set(x_16, 6, x_15); -x_17 = lean_io_get_num_heartbeats(x_3); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); +if (x_45 == 0) +{ +x_71 = x_61; +goto block_131; +} +else +{ +uint8_t x_133; +x_133 = 1; +x_71 = x_133; +goto block_131; +} +} +block_16: +{ +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +lean_dec(x_6); +lean_ctor_set(x_4, 0, x_7); +return x_4; +} +else +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_9); +lean_inc(x_8); +lean_dec(x_4); +x_10 = lean_ctor_get(x_8, 0); +lean_inc(x_10); +lean_dec(x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +else +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) +{ +return x_4; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} +} +} +block_38: +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_19; lean_object* x_20; x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_PPContext_runCoreM___rarg___closed__18; -x_21 = l_Lean_PPContext_runCoreM___rarg___closed__3; -x_22 = lean_unsigned_to_nat(0u); -x_23 = lean_unsigned_to_nat(1000u); -x_24 = lean_box(0); -x_25 = l_Lean_firstFrontendMacroScope; -x_26 = 0; -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_27 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_27, 0, x_20); -lean_ctor_set(x_27, 1, x_21); -lean_ctor_set(x_27, 2, x_5); -lean_ctor_set(x_27, 3, x_22); -lean_ctor_set(x_27, 4, x_23); -lean_ctor_set(x_27, 5, x_24); -lean_ctor_set(x_27, 6, x_6); -lean_ctor_set(x_27, 7, x_7); -lean_ctor_set(x_27, 8, x_18); -lean_ctor_set(x_27, 9, x_8); -lean_ctor_set(x_27, 10, x_25); -lean_ctor_set_uint8(x_27, sizeof(void*)*11, x_26); -lean_ctor_set_uint8(x_27, sizeof(void*)*11 + 1, x_10); -x_28 = lean_st_mk_ref(x_16, x_19); -x_29 = lean_ctor_get(x_28, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -lean_dec(x_28); -lean_inc(x_29); -x_31 = lean_apply_3(x_2, x_27, x_29, x_30); -if (lean_obj_tag(x_31) == 0) +x_20 = l_Lean_MessageData_toString(x_19, x_18); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -x_33 = lean_ctor_get(x_31, 1); -lean_inc(x_33); -lean_dec(x_31); -x_34 = lean_st_ref_get(x_29, x_33); -lean_dec(x_29); -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_32); -return x_34; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set_tag(x_20, 1); +lean_ctor_set(x_20, 0, x_23); +x_4 = x_20; +goto block_16; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_32); -lean_ctor_set(x_38, 1, x_37); -return x_38; +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_alloc_ctor(18, 1, 0); +lean_ctor_set(x_26, 0, x_24); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_4 = x_27; +goto block_16; } } else { -lean_object* x_39; -lean_dec(x_29); -x_39 = lean_ctor_get(x_31, 0); -lean_inc(x_39); -if (lean_obj_tag(x_39) == 0) +uint8_t x_28; +x_28 = !lean_is_exclusive(x_20); +if (x_28 == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_31, 1); -lean_inc(x_40); -lean_dec(x_31); -x_41 = lean_ctor_get(x_39, 1); -lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_MessageData_toString(x_41, x_40); -if (lean_obj_tag(x_42) == 0) +x_4 = x_20; +goto block_16; +} +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); +x_4 = x_31; +goto block_16; +} +} +} +else { -uint8_t x_43; -x_43 = !lean_is_exclusive(x_42); -if (x_43 == 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; +x_32 = lean_ctor_get(x_17, 0); +lean_inc(x_32); +lean_dec(x_17); +x_33 = l___private_Init_Data_Repr_0__Nat_reprFast(x_32); +x_34 = l_Lean_PPContext_runCoreM___rarg___closed__1; +x_35 = lean_string_append(x_34, x_33); +lean_dec(x_33); +x_36 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_36, 0, x_35); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_18); +x_4 = x_37; +goto block_16; +} +} +block_131: { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_42, 0); -x_45 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set_tag(x_42, 1); -lean_ctor_set(x_42, 0, x_45); -return x_42; +if (x_71 == 0) +{ +lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_72 = lean_st_ref_take(x_64, x_68); +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_is_exclusive(x_73); +if (x_75 == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_76 = lean_ctor_get(x_73, 0); +x_77 = lean_ctor_get(x_73, 4); +lean_dec(x_77); +x_78 = l_Lean_Kernel_enableDiag(x_76, x_45); +lean_ctor_set(x_73, 4, x_49); +lean_ctor_set(x_73, 0, x_78); +x_79 = lean_st_ref_set(x_64, x_73, x_74); +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = lean_box(0); +lean_inc(x_64); +x_82 = l_Lean_PPContext_runCoreM___rarg___lambda__1(x_40, x_45, x_2, x_81, x_62, x_64, x_80); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; uint8_t x_86; +x_83 = lean_ctor_get(x_82, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_82, 1); +lean_inc(x_84); +lean_dec(x_82); +x_85 = lean_st_ref_get(x_64, x_84); +lean_dec(x_64); +x_86 = !lean_is_exclusive(x_85); +if (x_86 == 0) +{ +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_85, 0); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_83); +lean_ctor_set(x_88, 1, x_87); +lean_ctor_set(x_85, 0, x_88); +x_4 = x_85; +goto block_16; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_46 = lean_ctor_get(x_42, 0); -x_47 = lean_ctor_get(x_42, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_42); -x_48 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_48, 0, x_46); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -return x_49; +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_89 = lean_ctor_get(x_85, 0); +x_90 = lean_ctor_get(x_85, 1); +lean_inc(x_90); +lean_inc(x_89); +lean_dec(x_85); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_83); +lean_ctor_set(x_91, 1, x_89); +x_92 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_90); +x_4 = x_92; +goto block_16; } } else { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_42); -if (x_50 == 0) +lean_object* x_93; lean_object* x_94; +lean_dec(x_64); +x_93 = lean_ctor_get(x_82, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_82, 1); +lean_inc(x_94); +lean_dec(x_82); +x_17 = x_93; +x_18 = x_94; +goto block_38; +} +} +else { -return x_42; +lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_95 = lean_ctor_get(x_73, 0); +x_96 = lean_ctor_get(x_73, 1); +x_97 = lean_ctor_get(x_73, 2); +x_98 = lean_ctor_get(x_73, 3); +x_99 = lean_ctor_get(x_73, 5); +x_100 = lean_ctor_get(x_73, 6); +lean_inc(x_100); +lean_inc(x_99); +lean_inc(x_98); +lean_inc(x_97); +lean_inc(x_96); +lean_inc(x_95); +lean_dec(x_73); +x_101 = l_Lean_Kernel_enableDiag(x_95, x_45); +x_102 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_102, 0, x_101); +lean_ctor_set(x_102, 1, x_96); +lean_ctor_set(x_102, 2, x_97); +lean_ctor_set(x_102, 3, x_98); +lean_ctor_set(x_102, 4, x_49); +lean_ctor_set(x_102, 5, x_99); +lean_ctor_set(x_102, 6, x_100); +x_103 = lean_st_ref_set(x_64, x_102, x_74); +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +lean_dec(x_103); +x_105 = lean_box(0); +lean_inc(x_64); +x_106 = l_Lean_PPContext_runCoreM___rarg___lambda__1(x_40, x_45, x_2, x_105, x_62, x_64, x_104); +if (lean_obj_tag(x_106) == 0) +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_107 = lean_ctor_get(x_106, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_106, 1); +lean_inc(x_108); +lean_dec(x_106); +x_109 = lean_st_ref_get(x_64, x_108); +lean_dec(x_64); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +if (lean_is_exclusive(x_109)) { + lean_ctor_release(x_109, 0); + lean_ctor_release(x_109, 1); + x_112 = x_109; +} else { + lean_dec_ref(x_109); + x_112 = lean_box(0); +} +x_113 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_113, 0, x_107); +lean_ctor_set(x_113, 1, x_110); +if (lean_is_scalar(x_112)) { + x_114 = lean_alloc_ctor(0, 2, 0); +} else { + x_114 = x_112; +} +lean_ctor_set(x_114, 0, x_113); +lean_ctor_set(x_114, 1, x_111); +x_4 = x_114; +goto block_16; } 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_115; lean_object* x_116; +lean_dec(x_64); +x_115 = lean_ctor_get(x_106, 0); +lean_inc(x_115); +x_116 = lean_ctor_get(x_106, 1); +lean_inc(x_116); +lean_dec(x_106); +x_17 = x_115; +x_18 = x_116; +goto block_38; } } } else { -uint8_t x_54; -x_54 = !lean_is_exclusive(x_31); -if (x_54 == 0) +lean_object* x_117; lean_object* x_118; +x_117 = lean_box(0); +lean_inc(x_64); +x_118 = l_Lean_PPContext_runCoreM___rarg___lambda__1(x_40, x_45, x_2, x_117, x_62, x_64, x_68); +if (lean_obj_tag(x_118) == 0) +{ +lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; +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 = lean_st_ref_get(x_64, x_120); +lean_dec(x_64); +x_122 = !lean_is_exclusive(x_121); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_121, 0); +x_124 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_124, 0, x_119); +lean_ctor_set(x_124, 1, x_123); +lean_ctor_set(x_121, 0, x_124); +x_4 = x_121; +goto block_16; +} +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; -x_55 = lean_ctor_get(x_31, 0); -lean_dec(x_55); -x_56 = lean_ctor_get(x_39, 0); -lean_inc(x_56); -lean_dec(x_39); -x_57 = l___private_Init_Data_Repr_0__Nat_reprFast(x_56); -x_58 = l_Lean_PPContext_runCoreM___rarg___closed__19; -x_59 = lean_string_append(x_58, x_57); -lean_dec(x_57); -x_60 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set(x_31, 0, x_60); -return x_31; +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +x_125 = lean_ctor_get(x_121, 0); +x_126 = lean_ctor_get(x_121, 1); +lean_inc(x_126); +lean_inc(x_125); +lean_dec(x_121); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_119); +lean_ctor_set(x_127, 1, x_125); +x_128 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_128, 0, x_127); +lean_ctor_set(x_128, 1, x_126); +x_4 = x_128; +goto block_16; +} } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_61 = lean_ctor_get(x_31, 1); -lean_inc(x_61); -lean_dec(x_31); -x_62 = lean_ctor_get(x_39, 0); -lean_inc(x_62); -lean_dec(x_39); -x_63 = l___private_Init_Data_Repr_0__Nat_reprFast(x_62); -x_64 = l_Lean_PPContext_runCoreM___rarg___closed__19; -x_65 = lean_string_append(x_64, x_63); -lean_dec(x_63); -x_66 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_66, 0, x_65); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_66); -lean_ctor_set(x_67, 1, x_61); -return x_67; +lean_object* x_129; lean_object* x_130; +lean_dec(x_64); +x_129 = lean_ctor_get(x_118, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_118, 1); +lean_inc(x_130); +lean_dec(x_118); +x_17 = x_129; +x_18 = x_130; +goto block_38; } } } @@ -630,17 +976,19 @@ LEAN_EXPORT lean_object* l_Lean_PPContext_runCoreM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PPContext_runCoreM___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PPContext_runCoreM___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PPContext_runCoreM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PPContext_runCoreM___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) { _start: { -lean_object* x_4; -x_4 = l_Lean_PPContext_runCoreM___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; +uint8_t x_8; lean_object* x_9; +x_8 = lean_unbox(x_2); +lean_dec(x_2); +x_9 = l_Lean_PPContext_runCoreM___rarg___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_4); +return x_9; } } static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__1() { @@ -672,7 +1020,7 @@ static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__13; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__14; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -692,7 +1040,7 @@ static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__13; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__14; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -764,7 +1112,7 @@ static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__13; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__14; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -798,7 +1146,7 @@ static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__13; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__14; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -842,7 +1190,7 @@ static lean_object* _init_l_Lean_PPContext_runMetaM___rarg___closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__14; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__15; x_2 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -855,12 +1203,12 @@ LEAN_EXPORT lean_object* l_Lean_PPContext_runMetaM___rarg(lean_object* x_1, lean { lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_4 = lean_ctor_get(x_1, 2); +lean_inc(x_4); x_5 = lean_box(0); x_6 = l_Lean_PPContext_runMetaM___rarg___closed__1; -x_7 = l_Lean_PPContext_runCoreM___rarg___closed__1; +x_7 = l_Lean_PPContext_runCoreM___rarg___closed__2; x_8 = lean_unsigned_to_nat(0u); x_9 = 0; -lean_inc(x_4); x_10 = lean_alloc_ctor(0, 6, 2); lean_ctor_set(x_10, 0, x_6); lean_ctor_set(x_10, 1, x_4); @@ -871,11 +1219,11 @@ lean_ctor_set(x_10, 5, x_5); lean_ctor_set_uint8(x_10, sizeof(void*)*6, x_9); lean_ctor_set_uint8(x_10, sizeof(void*)*6 + 1, x_9); x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); x_12 = lean_box(0); x_13 = l_Lean_PPContext_runMetaM___rarg___closed__16; -x_14 = l_Lean_PPContext_runCoreM___rarg___closed__11; +x_14 = l_Lean_PPContext_runCoreM___rarg___closed__12; x_15 = l_Lean_PPContext_runMetaM___rarg___closed__17; -lean_inc(x_11); x_16 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_16, 0, x_11); lean_ctor_set(x_16, 1, x_13); @@ -894,19 +1242,10 @@ LEAN_EXPORT lean_object* l_Lean_PPContext_runMetaM(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PPContext_runMetaM___rarg___boxed), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PPContext_runMetaM___rarg), 3, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PPContext_runMetaM___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_PPContext_runMetaM___rarg(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppCategory(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -1055,7 +1394,7 @@ lean_dec(x_12); x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_ppUsing___lambda__1), 7, 2); lean_closure_set(x_14, 0, x_2); lean_closure_set(x_14, 1, x_1); -x_15 = l_Lean_PPContext_runCoreM___rarg___closed__1; +x_15 = l_Lean_PPContext_runCoreM___rarg___closed__2; x_16 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_13, x_15, x_14, x_3, x_4, x_5, x_6, x_7); return x_16; } @@ -1164,321 +1503,1255 @@ return x_25; } else { -uint8_t x_26; -lean_dec(x_7); -lean_dec(x_6); -x_26 = !lean_is_exclusive(x_9); -if (x_26 == 0) +uint8_t x_26; +lean_dec(x_7); +lean_dec(x_6); +x_26 = !lean_is_exclusive(x_9); +if (x_26 == 0) +{ +return x_9; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_9, 0); +x_28 = lean_ctor_get(x_9, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_9); +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_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprWithInfos(lean_object* x_1, lean_object* x_2, 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; +x_9 = lean_ctor_get(x_4, 1); +lean_inc(x_9); +x_10 = lean_ctor_get(x_6, 2); +lean_inc(x_10); +x_11 = lean_box(0); +x_12 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_12, 0, x_10); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set(x_12, 2, x_11); +x_13 = l_Lean_LocalContext_sanitizeNames(x_9, x_12); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_ppExprWithInfos___lambda__1), 8, 3); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); +lean_closure_set(x_15, 2, x_3); +x_16 = l_Lean_PPContext_runCoreM___rarg___closed__2; +x_17 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_14, x_16, x_15, x_4, x_5, x_6, x_7, x_8); +return x_17; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprLegacy___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) { +_start: +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_11 = lean_ctor_get(x_7, 4); +lean_dec(x_11); +x_12 = lean_ctor_get(x_7, 2); +lean_dec(x_12); +x_13 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_14 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_13); +lean_ctor_set(x_7, 4, x_14); +lean_ctor_set(x_7, 2, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*11 + 1, x_2); +x_15 = l_Lean_PrettyPrinter_ppExpr___closed__1; +x_16 = l_Lean_PrettyPrinter_ppUsing(x_3, x_15, x_4, x_5, x_7, x_8, x_9); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_17 = lean_ctor_get(x_7, 0); +x_18 = lean_ctor_get(x_7, 1); +x_19 = lean_ctor_get(x_7, 3); +x_20 = lean_ctor_get(x_7, 5); +x_21 = lean_ctor_get(x_7, 6); +x_22 = lean_ctor_get(x_7, 7); +x_23 = lean_ctor_get(x_7, 8); +x_24 = lean_ctor_get(x_7, 9); +x_25 = lean_ctor_get(x_7, 10); +x_26 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +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_dec(x_7); +x_27 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_28 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_27); +x_29 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_29, 0, x_17); +lean_ctor_set(x_29, 1, x_18); +lean_ctor_set(x_29, 2, x_1); +lean_ctor_set(x_29, 3, x_19); +lean_ctor_set(x_29, 4, x_28); +lean_ctor_set(x_29, 5, x_20); +lean_ctor_set(x_29, 6, x_21); +lean_ctor_set(x_29, 7, x_22); +lean_ctor_set(x_29, 8, x_23); +lean_ctor_set(x_29, 9, x_24); +lean_ctor_set(x_29, 10, x_25); +lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*11 + 1, x_2); +x_30 = l_Lean_PrettyPrinter_ppExpr___closed__1; +x_31 = l_Lean_PrettyPrinter_ppUsing(x_3, x_30, x_4, x_5, x_29, x_8, x_9); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprLegacy___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; +x_13 = lean_ctor_get(x_9, 4); +lean_dec(x_13); +x_14 = lean_ctor_get(x_9, 2); +lean_dec(x_14); +x_15 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_16 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_15); +lean_ctor_set(x_9, 4, x_16); +lean_ctor_set(x_9, 2, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*11 + 1, x_2); +x_17 = lean_st_mk_ref(x_3, 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 = l_Lean_PPContext_runCoreM___rarg___closed__5; +x_21 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_20); +x_22 = lean_st_ref_get(x_10, x_19); +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 = l_Lean_Kernel_isDiagnosticsEnabled(x_25); +lean_dec(x_25); +if (x_26 == 0) +{ +if (x_21 == 0) +{ +uint8_t x_86; +x_86 = 1; +x_27 = x_86; +goto block_85; +} +else +{ +uint8_t x_87; +x_87 = 0; +x_27 = x_87; +goto block_85; +} +} +else +{ +if (x_21 == 0) +{ +uint8_t x_88; +x_88 = 0; +x_27 = x_88; +goto block_85; +} +else +{ +uint8_t x_89; +x_89 = 1; +x_27 = x_89; +goto block_85; +} +} +block_85: +{ +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = lean_st_ref_take(x_10, x_24); +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; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_32 = lean_ctor_get(x_29, 0); +x_33 = lean_ctor_get(x_29, 4); +lean_dec(x_33); +x_34 = l_Lean_Kernel_enableDiag(x_32, x_21); +lean_ctor_set(x_29, 4, x_7); +lean_ctor_set(x_29, 0, x_34); +x_35 = lean_st_ref_set(x_10, x_29, x_30); +x_36 = lean_ctor_get(x_35, 1); +lean_inc(x_36); +lean_dec(x_35); +x_37 = lean_box(0); +lean_inc(x_18); +x_38 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_4, x_21, x_5, x_6, x_18, x_37, x_9, x_10, x_36); +if (lean_obj_tag(x_38) == 0) +{ +lean_object* x_39; lean_object* x_40; lean_object* x_41; uint8_t x_42; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = lean_st_ref_get(x_18, x_40); +lean_dec(x_18); +x_42 = !lean_is_exclusive(x_41); +if (x_42 == 0) +{ +lean_object* x_43; +x_43 = lean_ctor_get(x_41, 0); +lean_dec(x_43); +lean_ctor_set(x_41, 0, x_39); +return x_41; +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_41, 1); +lean_inc(x_44); +lean_dec(x_41); +x_45 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_45, 0, x_39); +lean_ctor_set(x_45, 1, x_44); +return x_45; +} +} +else +{ +uint8_t x_46; +lean_dec(x_18); +x_46 = !lean_is_exclusive(x_38); +if (x_46 == 0) +{ +return x_38; +} +else +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_38, 0); +x_48 = lean_ctor_get(x_38, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_38); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +return x_49; +} +} +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_50 = lean_ctor_get(x_29, 0); +x_51 = lean_ctor_get(x_29, 1); +x_52 = lean_ctor_get(x_29, 2); +x_53 = lean_ctor_get(x_29, 3); +x_54 = lean_ctor_get(x_29, 5); +x_55 = lean_ctor_get(x_29, 6); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_29); +x_56 = l_Lean_Kernel_enableDiag(x_50, x_21); +x_57 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_51); +lean_ctor_set(x_57, 2, x_52); +lean_ctor_set(x_57, 3, x_53); +lean_ctor_set(x_57, 4, x_7); +lean_ctor_set(x_57, 5, x_54); +lean_ctor_set(x_57, 6, x_55); +x_58 = lean_st_ref_set(x_10, x_57, x_30); +x_59 = lean_ctor_get(x_58, 1); +lean_inc(x_59); +lean_dec(x_58); +x_60 = lean_box(0); +lean_inc(x_18); +x_61 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_4, x_21, x_5, x_6, x_18, x_60, x_9, x_10, x_59); +if (lean_obj_tag(x_61) == 0) +{ +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_62 = lean_ctor_get(x_61, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_61, 1); +lean_inc(x_63); +lean_dec(x_61); +x_64 = lean_st_ref_get(x_18, x_63); +lean_dec(x_18); +x_65 = lean_ctor_get(x_64, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_64)) { + lean_ctor_release(x_64, 0); + lean_ctor_release(x_64, 1); + x_66 = x_64; +} else { + lean_dec_ref(x_64); + x_66 = lean_box(0); +} +if (lean_is_scalar(x_66)) { + x_67 = lean_alloc_ctor(0, 2, 0); +} else { + x_67 = x_66; +} +lean_ctor_set(x_67, 0, x_62); +lean_ctor_set(x_67, 1, x_65); +return x_67; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; +lean_dec(x_18); +x_68 = lean_ctor_get(x_61, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_61, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_70 = x_61; +} else { + lean_dec_ref(x_61); + x_70 = lean_box(0); +} +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(1, 2, 0); +} else { + x_71 = x_70; +} +lean_ctor_set(x_71, 0, x_68); +lean_ctor_set(x_71, 1, x_69); +return x_71; +} +} +} +else +{ +lean_object* x_72; lean_object* x_73; +lean_dec(x_7); +x_72 = lean_box(0); +lean_inc(x_18); +x_73 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_4, x_21, x_5, x_6, x_18, x_72, x_9, x_10, x_24); +if (lean_obj_tag(x_73) == 0) +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +lean_dec(x_73); +x_76 = lean_st_ref_get(x_18, x_75); +lean_dec(x_18); +x_77 = !lean_is_exclusive(x_76); +if (x_77 == 0) +{ +lean_object* x_78; +x_78 = lean_ctor_get(x_76, 0); +lean_dec(x_78); +lean_ctor_set(x_76, 0, x_74); +return x_76; +} +else +{ +lean_object* x_79; lean_object* x_80; +x_79 = lean_ctor_get(x_76, 1); +lean_inc(x_79); +lean_dec(x_76); +x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_80, 0, x_74); +lean_ctor_set(x_80, 1, x_79); +return x_80; +} +} +else +{ +uint8_t x_81; +lean_dec(x_18); +x_81 = !lean_is_exclusive(x_73); +if (x_81 == 0) +{ +return x_73; +} +else +{ +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_73, 0); +x_83 = lean_ctor_get(x_73, 1); +lean_inc(x_83); +lean_inc(x_82); +lean_dec(x_73); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; +} +} +} +} +} +else +{ +lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; uint8_t x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; uint8_t x_113; +x_90 = lean_ctor_get(x_9, 0); +x_91 = lean_ctor_get(x_9, 1); +x_92 = lean_ctor_get(x_9, 3); +x_93 = lean_ctor_get(x_9, 5); +x_94 = lean_ctor_get(x_9, 6); +x_95 = lean_ctor_get(x_9, 7); +x_96 = lean_ctor_get(x_9, 8); +x_97 = lean_ctor_get(x_9, 9); +x_98 = lean_ctor_get(x_9, 10); +x_99 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); +lean_inc(x_98); +lean_inc(x_97); +lean_inc(x_96); +lean_inc(x_95); +lean_inc(x_94); +lean_inc(x_93); +lean_inc(x_92); +lean_inc(x_91); +lean_inc(x_90); +lean_dec(x_9); +x_100 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_101 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_100); +x_102 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_102, 0, x_90); +lean_ctor_set(x_102, 1, x_91); +lean_ctor_set(x_102, 2, x_1); +lean_ctor_set(x_102, 3, x_92); +lean_ctor_set(x_102, 4, x_101); +lean_ctor_set(x_102, 5, x_93); +lean_ctor_set(x_102, 6, x_94); +lean_ctor_set(x_102, 7, x_95); +lean_ctor_set(x_102, 8, x_96); +lean_ctor_set(x_102, 9, x_97); +lean_ctor_set(x_102, 10, x_98); +lean_ctor_set_uint8(x_102, sizeof(void*)*11, x_99); +lean_ctor_set_uint8(x_102, sizeof(void*)*11 + 1, x_2); +x_103 = lean_st_mk_ref(x_3, x_11); +x_104 = lean_ctor_get(x_103, 0); +lean_inc(x_104); +x_105 = lean_ctor_get(x_103, 1); +lean_inc(x_105); +lean_dec(x_103); +x_106 = l_Lean_PPContext_runCoreM___rarg___closed__5; +x_107 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_106); +x_108 = lean_st_ref_get(x_10, x_105); +x_109 = lean_ctor_get(x_108, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_108, 1); +lean_inc(x_110); +lean_dec(x_108); +x_111 = lean_ctor_get(x_109, 0); +lean_inc(x_111); +lean_dec(x_109); +x_112 = l_Lean_Kernel_isDiagnosticsEnabled(x_111); +lean_dec(x_111); +if (x_112 == 0) +{ +if (x_107 == 0) +{ +uint8_t x_153; +x_153 = 1; +x_113 = x_153; +goto block_152; +} +else +{ +uint8_t x_154; +x_154 = 0; +x_113 = x_154; +goto block_152; +} +} +else +{ +if (x_107 == 0) +{ +uint8_t x_155; +x_155 = 0; +x_113 = x_155; +goto block_152; +} +else +{ +uint8_t x_156; +x_156 = 1; +x_113 = x_156; +goto block_152; +} +} +block_152: +{ +if (x_113 == 0) +{ +lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; +x_114 = lean_st_ref_take(x_10, x_110); +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 = lean_ctor_get(x_115, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_115, 1); +lean_inc(x_118); +x_119 = lean_ctor_get(x_115, 2); +lean_inc(x_119); +x_120 = lean_ctor_get(x_115, 3); +lean_inc(x_120); +x_121 = lean_ctor_get(x_115, 5); +lean_inc(x_121); +x_122 = lean_ctor_get(x_115, 6); +lean_inc(x_122); +if (lean_is_exclusive(x_115)) { + lean_ctor_release(x_115, 0); + lean_ctor_release(x_115, 1); + lean_ctor_release(x_115, 2); + lean_ctor_release(x_115, 3); + lean_ctor_release(x_115, 4); + lean_ctor_release(x_115, 5); + lean_ctor_release(x_115, 6); + x_123 = x_115; +} else { + lean_dec_ref(x_115); + x_123 = lean_box(0); +} +x_124 = l_Lean_Kernel_enableDiag(x_117, x_107); +if (lean_is_scalar(x_123)) { + x_125 = lean_alloc_ctor(0, 7, 0); +} else { + x_125 = x_123; +} +lean_ctor_set(x_125, 0, x_124); +lean_ctor_set(x_125, 1, x_118); +lean_ctor_set(x_125, 2, x_119); +lean_ctor_set(x_125, 3, x_120); +lean_ctor_set(x_125, 4, x_7); +lean_ctor_set(x_125, 5, x_121); +lean_ctor_set(x_125, 6, x_122); +x_126 = lean_st_ref_set(x_10, x_125, x_116); +x_127 = lean_ctor_get(x_126, 1); +lean_inc(x_127); +lean_dec(x_126); +x_128 = lean_box(0); +lean_inc(x_104); +x_129 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_4, x_107, x_5, x_6, x_104, x_128, x_102, x_10, x_127); +if (lean_obj_tag(x_129) == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_130 = lean_ctor_get(x_129, 0); +lean_inc(x_130); +x_131 = lean_ctor_get(x_129, 1); +lean_inc(x_131); +lean_dec(x_129); +x_132 = lean_st_ref_get(x_104, x_131); +lean_dec(x_104); +x_133 = lean_ctor_get(x_132, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_132)) { + lean_ctor_release(x_132, 0); + lean_ctor_release(x_132, 1); + x_134 = x_132; +} else { + lean_dec_ref(x_132); + x_134 = lean_box(0); +} +if (lean_is_scalar(x_134)) { + x_135 = lean_alloc_ctor(0, 2, 0); +} else { + x_135 = x_134; +} +lean_ctor_set(x_135, 0, x_130); +lean_ctor_set(x_135, 1, x_133); +return x_135; +} +else +{ +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; +lean_dec(x_104); +x_136 = lean_ctor_get(x_129, 0); +lean_inc(x_136); +x_137 = lean_ctor_get(x_129, 1); +lean_inc(x_137); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + x_138 = x_129; +} else { + lean_dec_ref(x_129); + x_138 = lean_box(0); +} +if (lean_is_scalar(x_138)) { + x_139 = lean_alloc_ctor(1, 2, 0); +} else { + x_139 = x_138; +} +lean_ctor_set(x_139, 0, x_136); +lean_ctor_set(x_139, 1, x_137); +return x_139; +} +} +else +{ +lean_object* x_140; lean_object* x_141; +lean_dec(x_7); +x_140 = lean_box(0); +lean_inc(x_104); +x_141 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_4, x_107, x_5, x_6, x_104, x_140, x_102, x_10, x_110); +if (lean_obj_tag(x_141) == 0) +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_142 = lean_ctor_get(x_141, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_141, 1); +lean_inc(x_143); +lean_dec(x_141); +x_144 = lean_st_ref_get(x_104, x_143); +lean_dec(x_104); +x_145 = lean_ctor_get(x_144, 1); +lean_inc(x_145); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_146 = x_144; +} else { + lean_dec_ref(x_144); + x_146 = lean_box(0); +} +if (lean_is_scalar(x_146)) { + x_147 = lean_alloc_ctor(0, 2, 0); +} else { + x_147 = x_146; +} +lean_ctor_set(x_147, 0, x_142); +lean_ctor_set(x_147, 1, x_145); +return x_147; +} +else +{ +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; +lean_dec(x_104); +x_148 = lean_ctor_get(x_141, 0); +lean_inc(x_148); +x_149 = lean_ctor_get(x_141, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_141)) { + lean_ctor_release(x_141, 0); + lean_ctor_release(x_141, 1); + x_150 = x_141; +} else { + lean_dec_ref(x_141); + x_150 = lean_box(0); +} +if (lean_is_scalar(x_150)) { + x_151 = lean_alloc_ctor(1, 2, 0); +} else { + x_151 = x_150; +} +lean_ctor_set(x_151, 0, x_148); +lean_ctor_set(x_151, 1, x_149); +return x_151; +} +} +} +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_box(0); +x_2 = l_Lean_Core_getMaxHeartbeats(x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("_uniq", 5); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_ppExprLegacy___closed__2; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_ppExprLegacy___closed__3; +x_2 = lean_unsigned_to_nat(1u); +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 uint8_t _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; uint8_t x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PPContext_runCoreM___rarg___closed__5; +x_3 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* lean_pp_expr(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_20; lean_object* x_21; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; uint8_t x_78; +x_42 = lean_box(0); +x_43 = l_Lean_PPContext_runMetaM___rarg___closed__1; +x_44 = l_Lean_PPContext_runCoreM___rarg___closed__2; +x_45 = lean_unsigned_to_nat(0u); +x_46 = 0; +x_47 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_47, 0, x_43); +lean_ctor_set(x_47, 1, x_3); +lean_ctor_set(x_47, 2, x_44); +lean_ctor_set(x_47, 3, x_42); +lean_ctor_set(x_47, 4, x_45); +lean_ctor_set(x_47, 5, x_42); +lean_ctor_set_uint8(x_47, sizeof(void*)*6, x_46); +lean_ctor_set_uint8(x_47, sizeof(void*)*6 + 1, x_46); +x_48 = lean_box(0); +x_49 = l_Lean_PPContext_runMetaM___rarg___closed__16; +x_50 = l_Lean_PPContext_runCoreM___rarg___closed__12; +x_51 = l_Lean_PPContext_runMetaM___rarg___closed__17; +x_52 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_52, 0, x_2); +lean_ctor_set(x_52, 1, x_49); +lean_ctor_set(x_52, 2, x_48); +lean_ctor_set(x_52, 3, x_50); +lean_ctor_set(x_52, 4, x_51); +x_53 = lean_box(0); +x_54 = l_Lean_PPContext_runCoreM___rarg___closed__6; +x_55 = l_Lean_PrettyPrinter_ppExprLegacy___closed__4; +x_56 = l_Lean_PPContext_runCoreM___rarg___closed__16; +x_57 = l_Lean_PPContext_runCoreM___rarg___closed__18; +x_58 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_58, 0, x_1); +lean_ctor_set(x_58, 1, x_54); +lean_ctor_set(x_58, 2, x_55); +lean_ctor_set(x_58, 3, x_50); +lean_ctor_set(x_58, 4, x_56); +lean_ctor_set(x_58, 5, x_50); +lean_ctor_set(x_58, 6, x_57); +x_59 = lean_io_get_num_heartbeats(x_6); +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_PPContext_runCoreM___rarg___closed__19; +x_63 = l_Lean_PPContext_runCoreM___rarg___closed__4; +x_64 = lean_unsigned_to_nat(1000u); +x_65 = lean_box(0); +x_66 = lean_box(0); +x_67 = l_Lean_PrettyPrinter_ppExprLegacy___closed__1; +x_68 = l_Lean_firstFrontendMacroScope; +x_69 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_69, 0, x_62); +lean_ctor_set(x_69, 1, x_63); +lean_ctor_set(x_69, 2, x_53); +lean_ctor_set(x_69, 3, x_45); +lean_ctor_set(x_69, 4, x_64); +lean_ctor_set(x_69, 5, x_65); +lean_ctor_set(x_69, 6, x_66); +lean_ctor_set(x_69, 7, x_53); +lean_ctor_set(x_69, 8, x_60); +lean_ctor_set(x_69, 9, x_67); +lean_ctor_set(x_69, 10, x_68); +lean_ctor_set_uint8(x_69, sizeof(void*)*11, x_46); +lean_ctor_set_uint8(x_69, sizeof(void*)*11 + 1, x_46); +x_70 = lean_st_mk_ref(x_58, x_61); +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_st_ref_get(x_71, x_72); +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, 0); +lean_inc(x_76); +lean_dec(x_74); +x_77 = l_Lean_Kernel_isDiagnosticsEnabled(x_76); +lean_dec(x_76); +if (x_77 == 0) +{ +uint8_t x_142; +x_142 = l_Lean_PrettyPrinter_ppExprLegacy___closed__5; +if (x_142 == 0) +{ +uint8_t x_143; +x_143 = 1; +x_78 = x_143; +goto block_141; +} +else +{ +x_78 = x_46; +goto block_141; +} +} +else +{ +uint8_t x_144; +x_144 = l_Lean_PrettyPrinter_ppExprLegacy___closed__5; +if (x_144 == 0) +{ +x_78 = x_46; +goto block_141; +} +else +{ +uint8_t x_145; +x_145 = 1; +x_78 = x_145; +goto block_141; +} +} +block_19: +{ +if (lean_obj_tag(x_7) == 0) +{ +uint8_t x_8; +x_8 = !lean_is_exclusive(x_7); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_7, 0); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +lean_dec(x_9); +lean_ctor_set(x_7, 0, x_10); +return x_7; +} +else +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_7, 0); +x_12 = lean_ctor_get(x_7, 1); +lean_inc(x_12); +lean_inc(x_11); +lean_dec(x_7); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_13); +lean_ctor_set(x_14, 1, x_12); +return x_14; +} +} +else +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_7); +if (x_15 == 0) +{ +return x_7; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_7, 0); +x_17 = lean_ctor_get(x_7, 1); +lean_inc(x_17); +lean_inc(x_16); +lean_dec(x_7); +x_18 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +return x_18; +} +} +} +block_41: +{ +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_MessageData_toString(x_22, x_21); +if (lean_obj_tag(x_23) == 0) +{ +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -return x_9; +lean_object* x_25; lean_object* x_26; +x_25 = lean_ctor_get(x_23, 0); +x_26 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set_tag(x_23, 1); +lean_ctor_set(x_23, 0, x_26); +x_7 = x_23; +goto block_19; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_9, 0); -x_28 = lean_ctor_get(x_9, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_27 = lean_ctor_get(x_23, 0); +x_28 = lean_ctor_get(x_23, 1); lean_inc(x_28); lean_inc(x_27); -lean_dec(x_9); -x_29 = lean_alloc_ctor(1, 2, 0); +lean_dec(x_23); +x_29 = lean_alloc_ctor(18, 1, 0); lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -return x_29; -} +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_28); +x_7 = x_30; +goto block_19; } } +else +{ +uint8_t x_31; +x_31 = !lean_is_exclusive(x_23); +if (x_31 == 0) +{ +x_7 = x_23; +goto block_19; } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprWithInfos(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_9 = lean_ctor_get(x_4, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_6, 2); -lean_inc(x_10); -x_11 = lean_box(0); -x_12 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_12, 0, x_10); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set(x_12, 2, x_11); -x_13 = l_Lean_LocalContext_sanitizeNames(x_9, x_12); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -lean_dec(x_13); -x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_ppExprWithInfos___lambda__1), 8, 3); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -lean_closure_set(x_15, 2, x_3); -x_16 = l_Lean_PPContext_runCoreM___rarg___closed__1; -x_17 = l_Lean_Meta_withLCtx___at___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore___spec__2___rarg(x_14, x_16, x_15, x_4, x_5, x_6, x_7, x_8); -return x_17; +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_23, 0); +x_33 = lean_ctor_get(x_23, 1); +lean_inc(x_33); +lean_inc(x_32); +lean_dec(x_23); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_7 = x_34; +goto block_19; } } -static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("_uniq", 5); -return x_1; } +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; +x_35 = lean_ctor_get(x_20, 0); +lean_inc(x_35); +lean_dec(x_20); +x_36 = l___private_Init_Data_Repr_0__Nat_reprFast(x_35); +x_37 = l_Lean_PPContext_runCoreM___rarg___closed__1; +x_38 = lean_string_append(x_37, x_36); +lean_dec(x_36); +x_39 = lean_alloc_ctor(18, 1, 0); +lean_ctor_set(x_39, 0, x_38); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_21); +x_7 = x_40; +goto block_19; +} +} +block_141: +{ +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; +x_79 = lean_st_ref_take(x_71, x_75); +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_is_exclusive(x_80); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_83 = lean_ctor_get(x_80, 0); +x_84 = lean_ctor_get(x_80, 4); +lean_dec(x_84); +x_85 = l_Lean_PrettyPrinter_ppExprLegacy___closed__5; +x_86 = l_Lean_Kernel_enableDiag(x_83, x_85); +lean_ctor_set(x_80, 4, x_56); +lean_ctor_set(x_80, 0, x_86); +x_87 = lean_st_ref_set(x_71, x_80, x_81); +x_88 = lean_ctor_get(x_87, 1); +lean_inc(x_88); +lean_dec(x_87); +x_89 = lean_box(0); +lean_inc(x_71); +x_90 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__2(x_53, x_85, x_52, x_4, x_5, x_47, x_56, x_89, x_69, x_71, x_88); +if (lean_obj_tag(x_90) == 0) +{ +lean_object* x_91; lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); +lean_inc(x_92); +lean_dec(x_90); +x_93 = lean_st_ref_get(x_71, x_92); +lean_dec(x_71); +x_94 = !lean_is_exclusive(x_93); +if (x_94 == 0) +{ +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_93, 0); +x_96 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_96, 0, x_91); +lean_ctor_set(x_96, 1, x_95); +lean_ctor_set(x_93, 0, x_96); +x_7 = x_93; +goto block_19; } -static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_ppExprLegacy___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_97 = lean_ctor_get(x_93, 0); +x_98 = lean_ctor_get(x_93, 1); +lean_inc(x_98); +lean_inc(x_97); +lean_dec(x_93); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_91); +lean_ctor_set(x_99, 1, x_97); +x_100 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_100, 0, x_99); +lean_ctor_set(x_100, 1, x_98); +x_7 = x_100; +goto block_19; } } -static lean_object* _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__3() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_ppExprLegacy___closed__2; -x_2 = lean_unsigned_to_nat(1u); -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_101; lean_object* x_102; +lean_dec(x_71); +x_101 = lean_ctor_get(x_90, 0); +lean_inc(x_101); +x_102 = lean_ctor_get(x_90, 1); +lean_inc(x_102); +lean_dec(x_90); +x_20 = x_101; +x_21 = x_102; +goto block_41; } } -LEAN_EXPORT lean_object* lean_pp_expr(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; 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; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_7 = lean_box(0); -x_8 = l_Lean_PPContext_runMetaM___rarg___closed__1; -x_9 = l_Lean_PPContext_runCoreM___rarg___closed__1; -x_10 = lean_unsigned_to_nat(0u); -x_11 = 0; -x_12 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_12, 0, x_8); -lean_ctor_set(x_12, 1, x_3); -lean_ctor_set(x_12, 2, x_9); -lean_ctor_set(x_12, 3, x_7); -lean_ctor_set(x_12, 4, x_10); -lean_ctor_set(x_12, 5, x_7); -lean_ctor_set_uint8(x_12, sizeof(void*)*6, x_11); -lean_ctor_set_uint8(x_12, sizeof(void*)*6 + 1, x_11); -x_13 = lean_box(0); -x_14 = l_Lean_PPContext_runMetaM___rarg___closed__16; -x_15 = l_Lean_PPContext_runCoreM___rarg___closed__11; -x_16 = l_Lean_PPContext_runMetaM___rarg___closed__17; -x_17 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_14); -lean_ctor_set(x_17, 2, x_13); -lean_ctor_set(x_17, 3, x_15); -lean_ctor_set(x_17, 4, x_16); -x_18 = lean_box(0); -x_19 = l_Lean_Core_getMaxHeartbeats(x_4); -x_20 = l_Lean_PPContext_runCoreM___rarg___closed__4; -x_21 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_20); -x_22 = l_Lean_PPContext_runCoreM___rarg___closed__5; -x_23 = l_Lean_PrettyPrinter_ppExprLegacy___closed__3; -x_24 = l_Lean_PPContext_runCoreM___rarg___closed__15; -x_25 = l_Lean_PPContext_runCoreM___rarg___closed__17; -x_26 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_26, 0, x_1); -lean_ctor_set(x_26, 1, x_22); -lean_ctor_set(x_26, 2, x_23); -lean_ctor_set(x_26, 3, x_15); -lean_ctor_set(x_26, 4, x_24); -lean_ctor_set(x_26, 5, x_15); -lean_ctor_set(x_26, 6, x_25); -x_27 = lean_io_get_num_heartbeats(x_6); -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_PPContext_runCoreM___rarg___closed__18; -x_31 = l_Lean_PPContext_runCoreM___rarg___closed__3; -x_32 = lean_unsigned_to_nat(1000u); -x_33 = lean_box(0); -x_34 = lean_box(0); -x_35 = l_Lean_firstFrontendMacroScope; -x_36 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_36, 0, x_30); -lean_ctor_set(x_36, 1, x_31); -lean_ctor_set(x_36, 2, x_4); -lean_ctor_set(x_36, 3, x_10); -lean_ctor_set(x_36, 4, x_32); -lean_ctor_set(x_36, 5, x_33); -lean_ctor_set(x_36, 6, x_34); -lean_ctor_set(x_36, 7, x_18); -lean_ctor_set(x_36, 8, x_28); -lean_ctor_set(x_36, 9, x_19); -lean_ctor_set(x_36, 10, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*11, x_11); -lean_ctor_set_uint8(x_36, sizeof(void*)*11 + 1, x_21); -x_37 = lean_st_mk_ref(x_26, x_29); -x_38 = lean_ctor_get(x_37, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_st_mk_ref(x_17, x_39); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_PrettyPrinter_ppExpr___closed__1; -lean_inc(x_38); -lean_inc(x_41); -x_44 = l_Lean_PrettyPrinter_ppUsing(x_5, x_43, x_12, x_41, x_36, x_38, x_42); -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; uint8_t x_50; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_st_ref_get(x_41, x_46); -lean_dec(x_41); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); -x_49 = lean_st_ref_get(x_38, x_48); -lean_dec(x_38); -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +else { -lean_object* x_51; -x_51 = lean_ctor_get(x_49, 0); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_45); -return x_49; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_103 = lean_ctor_get(x_80, 0); +x_104 = lean_ctor_get(x_80, 1); +x_105 = lean_ctor_get(x_80, 2); +x_106 = lean_ctor_get(x_80, 3); +x_107 = lean_ctor_get(x_80, 5); +x_108 = lean_ctor_get(x_80, 6); +lean_inc(x_108); +lean_inc(x_107); +lean_inc(x_106); +lean_inc(x_105); +lean_inc(x_104); +lean_inc(x_103); +lean_dec(x_80); +x_109 = l_Lean_PrettyPrinter_ppExprLegacy___closed__5; +x_110 = l_Lean_Kernel_enableDiag(x_103, x_109); +x_111 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_111, 0, x_110); +lean_ctor_set(x_111, 1, x_104); +lean_ctor_set(x_111, 2, x_105); +lean_ctor_set(x_111, 3, x_106); +lean_ctor_set(x_111, 4, x_56); +lean_ctor_set(x_111, 5, x_107); +lean_ctor_set(x_111, 6, x_108); +x_112 = lean_st_ref_set(x_71, x_111, x_81); +x_113 = lean_ctor_get(x_112, 1); +lean_inc(x_113); +lean_dec(x_112); +x_114 = lean_box(0); +lean_inc(x_71); +x_115 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__2(x_53, x_109, x_52, x_4, x_5, x_47, x_56, x_114, x_69, x_71, x_113); +if (lean_obj_tag(x_115) == 0) +{ +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_116 = lean_ctor_get(x_115, 0); +lean_inc(x_116); +x_117 = lean_ctor_get(x_115, 1); +lean_inc(x_117); +lean_dec(x_115); +x_118 = lean_st_ref_get(x_71, x_117); +lean_dec(x_71); +x_119 = lean_ctor_get(x_118, 0); +lean_inc(x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_121 = x_118; +} else { + lean_dec_ref(x_118); + x_121 = lean_box(0); +} +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_116); +lean_ctor_set(x_122, 1, x_119); +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); +x_7 = x_123; +goto block_19; } else { -lean_object* x_52; lean_object* x_53; -x_52 = lean_ctor_get(x_49, 1); -lean_inc(x_52); -lean_dec(x_49); -x_53 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_53, 0, x_45); -lean_ctor_set(x_53, 1, x_52); -return x_53; +lean_object* x_124; lean_object* x_125; +lean_dec(x_71); +x_124 = lean_ctor_get(x_115, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_115, 1); +lean_inc(x_125); +lean_dec(x_115); +x_20 = x_124; +x_21 = x_125; +goto block_41; +} } } else { -lean_object* x_54; -lean_dec(x_41); -lean_dec(x_38); -x_54 = lean_ctor_get(x_44, 0); -lean_inc(x_54); -if (lean_obj_tag(x_54) == 0) -{ -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_44, 1); -lean_inc(x_55); -lean_dec(x_44); -x_56 = lean_ctor_get(x_54, 1); -lean_inc(x_56); -lean_dec(x_54); -x_57 = l_Lean_MessageData_toString(x_56, x_55); -if (lean_obj_tag(x_57) == 0) -{ -uint8_t x_58; -x_58 = !lean_is_exclusive(x_57); -if (x_58 == 0) -{ -lean_object* x_59; lean_object* x_60; -x_59 = lean_ctor_get(x_57, 0); -x_60 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_60, 0, x_59); -lean_ctor_set_tag(x_57, 1); -lean_ctor_set(x_57, 0, x_60); -return x_57; +uint8_t x_126; lean_object* x_127; lean_object* x_128; +x_126 = l_Lean_PrettyPrinter_ppExprLegacy___closed__5; +x_127 = lean_box(0); +lean_inc(x_71); +x_128 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__2(x_53, x_126, x_52, x_4, x_5, x_47, x_56, x_127, x_69, x_71, x_75); +if (lean_obj_tag(x_128) == 0) +{ +lean_object* x_129; lean_object* x_130; lean_object* x_131; uint8_t x_132; +x_129 = lean_ctor_get(x_128, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_128, 1); +lean_inc(x_130); +lean_dec(x_128); +x_131 = lean_st_ref_get(x_71, x_130); +lean_dec(x_71); +x_132 = !lean_is_exclusive(x_131); +if (x_132 == 0) +{ +lean_object* x_133; lean_object* x_134; +x_133 = lean_ctor_get(x_131, 0); +x_134 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_134, 0, x_129); +lean_ctor_set(x_134, 1, x_133); +lean_ctor_set(x_131, 0, x_134); +x_7 = x_131; +goto block_19; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_61 = lean_ctor_get(x_57, 0); -x_62 = lean_ctor_get(x_57, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_57); -x_63 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_63, 0, x_61); -x_64 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_62); -return x_64; +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_135 = lean_ctor_get(x_131, 0); +x_136 = lean_ctor_get(x_131, 1); +lean_inc(x_136); +lean_inc(x_135); +lean_dec(x_131); +x_137 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_137, 0, x_129); +lean_ctor_set(x_137, 1, x_135); +x_138 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_138, 0, x_137); +lean_ctor_set(x_138, 1, x_136); +x_7 = x_138; +goto block_19; } } else { -uint8_t x_65; -x_65 = !lean_is_exclusive(x_57); -if (x_65 == 0) -{ -return x_57; +lean_object* x_139; lean_object* x_140; +lean_dec(x_71); +x_139 = lean_ctor_get(x_128, 0); +lean_inc(x_139); +x_140 = lean_ctor_get(x_128, 1); +lean_inc(x_140); +lean_dec(x_128); +x_20 = x_139; +x_21 = x_140; +goto block_41; } -else -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_66 = lean_ctor_get(x_57, 0); -x_67 = lean_ctor_get(x_57, 1); -lean_inc(x_67); -lean_inc(x_66); -lean_dec(x_57); -x_68 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_68, 0, x_66); -lean_ctor_set(x_68, 1, x_67); -return x_68; } } } -else -{ -uint8_t x_69; -x_69 = !lean_is_exclusive(x_44); -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 = lean_ctor_get(x_44, 0); -lean_dec(x_70); -x_71 = lean_ctor_get(x_54, 0); -lean_inc(x_71); -lean_dec(x_54); -x_72 = l___private_Init_Data_Repr_0__Nat_reprFast(x_71); -x_73 = l_Lean_PPContext_runCoreM___rarg___closed__19; -x_74 = lean_string_append(x_73, x_72); -lean_dec(x_72); -x_75 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_44, 0, x_75); -return x_44; } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprLegacy___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_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_76 = lean_ctor_get(x_44, 1); -lean_inc(x_76); -lean_dec(x_44); -x_77 = lean_ctor_get(x_54, 0); -lean_inc(x_77); -lean_dec(x_54); -x_78 = l___private_Init_Data_Repr_0__Nat_reprFast(x_77); -x_79 = l_Lean_PPContext_runCoreM___rarg___closed__19; -x_80 = lean_string_append(x_79, x_78); -lean_dec(x_78); -x_81 = lean_alloc_ctor(18, 1, 0); -lean_ctor_set(x_81, 0, x_80); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_81); -lean_ctor_set(x_82, 1, x_76); -return x_82; -} +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_11; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_ppExprLegacy___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_2); +lean_dec(x_2); +x_13 = l_Lean_PrettyPrinter_ppExprLegacy___lambda__2(x_1, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_8); +return x_13; } } static lean_object* _init_l_Lean_PrettyPrinter_ppTactic___closed__1() { @@ -2059,7 +3332,7 @@ x_3 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPr return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -2267,7 +3540,7 @@ return x_44; } } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; @@ -2475,7 +3748,7 @@ return x_42; } } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; @@ -2683,7 +3956,7 @@ return x_44; } } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -2691,35 +3964,35 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delab), 7, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____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; lean_object* x_7; lean_object* x_8; x_4 = lean_box(0); -x_5 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___closed__1; +x_5 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1___closed__1; x_6 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_ppExprWithInfos), 8, 3); lean_closure_set(x_6, 0, x_2); lean_closure_set(x_6, 1, x_4); lean_closure_set(x_6, 2, x_5); -x_7 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__1), 6, 1); +x_7 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__1), 6, 1); lean_closure_set(x_7, 0, x_6); x_8 = l_Lean_PPContext_runMetaM___rarg(x_1, x_7, x_3); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__2(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_alloc_closure((void*)(l_Lean_PrettyPrinter_ppTerm), 4, 1); lean_closure_set(x_4, 0, x_2); -x_5 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__2), 4, 1); +x_5 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__2), 4, 1); lean_closure_set(x_5, 0, x_4); x_6 = l_Lean_PPContext_runCoreM___rarg(x_1, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; uint8_t x_5; lean_object* x_6; lean_object* x_7; @@ -2732,58 +4005,58 @@ lean_ctor_set(x_7, 1, x_3); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_alloc_closure((void*)(l_Lean_Meta_ppGoal), 6, 1); lean_closure_set(x_4, 0, x_2); -x_5 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____spec__3), 6, 1); +x_5 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_0__Lean_PrettyPrinter_withoutContext___at_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____spec__3), 6, 1); lean_closure_set(x_5, 0, x_4); x_6 = l_Lean_PPContext_runMetaM___rarg(x_1, x_5, x_3); return x_6; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__2___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__2), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__3___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__3___boxed), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__4___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__4), 3, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____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_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__1; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__2; -x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__3; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__4; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__1; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__2; +x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__3; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__4; x_5 = lean_alloc_ctor(0, 4, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_2); @@ -2792,7 +4065,7 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__6() { _start: { lean_object* x_1; @@ -2800,12 +4073,12 @@ x_1 = l_Lean_ppFnsRef; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__6; -x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__5; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__6; +x_3 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__5; x_4 = lean_st_ref_set(x_2, x_3, x_1); x_5 = !lean_is_exclusive(x_4); if (x_5 == 0) @@ -2827,43 +4100,16 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__2(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__3(x_1, x_2, x_3); -lean_dec(x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__4(x_1, x_2, x_3); +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__3(x_1, x_2, x_3); lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1() { _start: { lean_object* x_1; @@ -2871,17 +4117,17 @@ x_1 = lean_mk_string_from_bytes("PrettyPrinter", 13); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3() { _start: { lean_object* x_1; @@ -2889,27 +4135,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__4; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__4; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__6() { _start: { lean_object* x_1; @@ -2917,17 +4163,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__7() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__5; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__6; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__5; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__8() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__8() { _start: { lean_object* x_1; @@ -2935,37 +4181,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__9() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__7; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__8; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__7; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__10() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__9; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__9; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__10; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__10; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__12() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__12() { _start: { lean_object* x_1; @@ -2973,33 +4219,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__13() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__11; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__12; +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__11; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__14() { +static lean_object* _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__13; -x_2 = lean_unsigned_to_nat(807u); +x_1 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__13; +x_2 = lean_unsigned_to_nat(814u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__2; +x_2 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__2; x_3 = 0; -x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__14; +x_4 = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -3114,46 +4360,40 @@ x_6 = l_Lean_ParserCompiler_registerParserCompiler___rarg(x_5, x_4); return x_6; } } -static lean_object* _init_l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_pp_tagAppFns; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___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) { _start: { -uint8_t x_7; -x_7 = !lean_is_exclusive(x_4); -if (x_7 == 0) +uint8_t x_10; +x_10 = !lean_is_exclusive(x_7); +if (x_10 == 0) { -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; lean_object* x_12; -x_8 = lean_ctor_get(x_4, 2); -x_9 = l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1; -x_10 = 1; -x_11 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_8, x_9, x_10); -lean_ctor_set(x_4, 2, x_11); -x_12 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -return x_12; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_ctor_get(x_7, 4); +lean_dec(x_11); +x_12 = lean_ctor_get(x_7, 2); +lean_dec(x_12); +x_13 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_14 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_13); +lean_ctor_set(x_7, 4, x_14); +lean_ctor_set(x_7, 2, x_1); +lean_ctor_set_uint8(x_7, sizeof(void*)*11 + 1, x_2); +x_15 = lean_apply_5(x_3, x_4, x_5, x_7, x_8, x_9); +return x_15; } 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; uint8_t 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_13 = lean_ctor_get(x_4, 0); -x_14 = lean_ctor_get(x_4, 1); -x_15 = lean_ctor_get(x_4, 2); -x_16 = lean_ctor_get(x_4, 3); -x_17 = lean_ctor_get(x_4, 4); -x_18 = lean_ctor_get(x_4, 5); -x_19 = lean_ctor_get(x_4, 6); -x_20 = lean_ctor_get(x_4, 7); -x_21 = lean_ctor_get(x_4, 8); -x_22 = lean_ctor_get(x_4, 9); -x_23 = lean_ctor_get(x_4, 10); -x_24 = lean_ctor_get_uint8(x_4, sizeof(void*)*11); -x_25 = lean_ctor_get_uint8(x_4, sizeof(void*)*11 + 1); +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_object* x_28; lean_object* x_29; +x_16 = lean_ctor_get(x_7, 0); +x_17 = lean_ctor_get(x_7, 1); +x_18 = lean_ctor_get(x_7, 3); +x_19 = lean_ctor_get(x_7, 5); +x_20 = lean_ctor_get(x_7, 6); +x_21 = lean_ctor_get(x_7, 7); +x_22 = lean_ctor_get(x_7, 8); +x_23 = lean_ctor_get(x_7, 9); +x_24 = lean_ctor_get(x_7, 10); +x_25 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); +lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); lean_inc(x_21); @@ -3162,33 +4402,164 @@ lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); +lean_dec(x_7); +x_26 = l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1; +x_27 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_26); +x_28 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_28, 0, x_16); +lean_ctor_set(x_28, 1, x_17); +lean_ctor_set(x_28, 2, x_1); +lean_ctor_set(x_28, 3, x_18); +lean_ctor_set(x_28, 4, x_27); +lean_ctor_set(x_28, 5, x_19); +lean_ctor_set(x_28, 6, x_20); +lean_ctor_set(x_28, 7, x_21); +lean_ctor_set(x_28, 8, x_22); +lean_ctor_set(x_28, 9, x_23); +lean_ctor_set(x_28, 10, x_24); +lean_ctor_set_uint8(x_28, sizeof(void*)*11, x_25); +lean_ctor_set_uint8(x_28, sizeof(void*)*11 + 1, x_2); +x_29 = lean_apply_5(x_3, x_4, x_5, x_28, x_8, x_9); +return x_29; +} +} +} +static lean_object* _init_l_Lean_MessageData_ofFormatWithInfos___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_tagAppFns; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___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; 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; uint8_t x_17; uint8_t x_18; +x_7 = lean_ctor_get(x_4, 2); +lean_inc(x_7); +x_8 = l_Lean_MessageData_ofFormatWithInfos___lambda__2___closed__1; +x_9 = 1; +x_10 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_7, x_8, x_9); +x_11 = l_Lean_PPContext_runCoreM___rarg___closed__5; +x_12 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_10, x_11); +x_13 = lean_st_ref_get(x_5, x_6); +x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_4); -x_26 = l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1; -x_27 = 1; -x_28 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_15, x_26, x_27); -x_29 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_29, 0, x_13); -lean_ctor_set(x_29, 1, x_14); -lean_ctor_set(x_29, 2, x_28); -lean_ctor_set(x_29, 3, x_16); -lean_ctor_set(x_29, 4, x_17); -lean_ctor_set(x_29, 5, x_18); -lean_ctor_set(x_29, 6, x_19); -lean_ctor_set(x_29, 7, x_20); -lean_ctor_set(x_29, 8, x_21); -lean_ctor_set(x_29, 9, x_22); -lean_ctor_set(x_29, 10, x_23); -lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_24); -lean_ctor_set_uint8(x_29, sizeof(void*)*11 + 1, x_25); -x_30 = lean_apply_5(x_1, x_2, x_3, x_29, x_5, x_6); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Kernel_isDiagnosticsEnabled(x_16); +lean_dec(x_16); +if (x_17 == 0) +{ +if (x_12 == 0) +{ +x_18 = x_9; +goto block_46; +} +else +{ +uint8_t x_47; +x_47 = 0; +x_18 = x_47; +goto block_46; +} +} +else +{ +if (x_12 == 0) +{ +uint8_t x_48; +x_48 = 0; +x_18 = x_48; +goto block_46; +} +else +{ +x_18 = x_9; +goto block_46; +} +} +block_46: +{ +if (x_18 == 0) +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_st_ref_take(x_5, x_15); +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = !lean_is_exclusive(x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_20, 0); +x_24 = lean_ctor_get(x_20, 4); +lean_dec(x_24); +x_25 = l_Lean_Kernel_enableDiag(x_23, x_12); +x_26 = l_Lean_PPContext_runCoreM___rarg___closed__16; +lean_ctor_set(x_20, 4, x_26); +lean_ctor_set(x_20, 0, x_25); +x_27 = lean_st_ref_set(x_5, x_20, x_21); +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_box(0); +x_30 = l_Lean_MessageData_ofFormatWithInfos___lambda__1(x_10, x_12, x_1, x_2, x_3, x_29, x_4, x_5, 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; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +x_33 = lean_ctor_get(x_20, 2); +x_34 = lean_ctor_get(x_20, 3); +x_35 = lean_ctor_get(x_20, 5); +x_36 = lean_ctor_get(x_20, 6); +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_20); +x_37 = l_Lean_Kernel_enableDiag(x_31, x_12); +x_38 = l_Lean_PPContext_runCoreM___rarg___closed__16; +x_39 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_32); +lean_ctor_set(x_39, 2, x_33); +lean_ctor_set(x_39, 3, x_34); +lean_ctor_set(x_39, 4, x_38); +lean_ctor_set(x_39, 5, x_35); +lean_ctor_set(x_39, 6, x_36); +x_40 = lean_st_ref_set(x_5, x_39, x_21); +x_41 = lean_ctor_get(x_40, 1); +lean_inc(x_41); +lean_dec(x_40); +x_42 = lean_box(0); +x_43 = l_Lean_MessageData_ofFormatWithInfos___lambda__1(x_10, x_12, x_1, x_2, x_3, x_42, x_4, x_5, x_41); +return x_43; +} +} +else +{ +lean_object* x_44; lean_object* x_45; +x_44 = lean_box(0); +x_45 = l_Lean_MessageData_ofFormatWithInfos___lambda__1(x_10, x_12, x_1, x_2, x_3, x_44, x_4, x_5, x_15); +return x_45; +} +} } } -LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_3) == 0) @@ -3211,14 +4582,16 @@ else lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_dec(x_1); x_10 = lean_ctor_get(x_3, 0); -x_11 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__1), 6, 1); +lean_inc(x_10); +lean_dec(x_3); +x_11 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__2), 6, 1); lean_closure_set(x_11, 0, x_2); x_12 = l_Lean_PPContext_runMetaM___rarg(x_10, x_11, x_4); return x_12; } } } -LEAN_EXPORT uint8_t l_Lean_MessageData_ofFormatWithInfos___lambda__3(lean_object* x_1) { +LEAN_EXPORT uint8_t l_Lean_MessageData_ofFormatWithInfos___lambda__4(lean_object* x_1) { _start: { uint8_t x_2; @@ -3230,7 +4603,7 @@ static lean_object* _init_l_Lean_MessageData_ofFormatWithInfos___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__3___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__4___boxed), 1, 0); return x_1; } } @@ -3238,7 +4611,7 @@ LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos(lean_object* x_1, _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_3 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__2___boxed), 4, 2); +x_3 = lean_alloc_closure((void*)(l_Lean_MessageData_ofFormatWithInfos___lambda__3), 4, 2); lean_closure_set(x_3, 0, x_2); lean_closure_set(x_3, 1, x_1); x_4 = l_Lean_MessageData_ofFormatWithInfos___closed__1; @@ -3250,20 +4623,22 @@ lean_ctor_set(x_6, 0, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___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_MessageData_ofFormatWithInfos___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_5; -x_5 = l_Lean_MessageData_ofFormatWithInfos___lambda__2(x_1, x_2, x_3, x_4); -lean_dec(x_3); -return x_5; +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l_Lean_MessageData_ofFormatWithInfos___lambda__1(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_6); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___lambda__3___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_MessageData_ofFormatWithInfos___lambda__4___boxed(lean_object* x_1) { _start: { uint8_t x_2; lean_object* x_3; -x_2 = l_Lean_MessageData_ofFormatWithInfos___lambda__3(x_1); +x_2 = l_Lean_MessageData_ofFormatWithInfos___lambda__4(x_1); lean_dec(x_1); x_3 = lean_box(x_2); return x_3; @@ -3282,7 +4657,7 @@ static lean_object* _init_l_Lean_MessageData_ofConst___lambda__1___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PPContext_runCoreM___rarg___closed__2; +x_1 = l_Lean_PPContext_runCoreM___rarg___closed__3; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; @@ -3455,6 +4830,8 @@ lean_dec_ref(res); res = initialize_Lean_ParserCompiler(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1 = _init_l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_PPContext_runCoreM___rarg___lambda__1___closed__1); l_Lean_PPContext_runCoreM___rarg___closed__1 = _init_l_Lean_PPContext_runCoreM___rarg___closed__1(); lean_mark_persistent(l_Lean_PPContext_runCoreM___rarg___closed__1); l_Lean_PPContext_runCoreM___rarg___closed__2 = _init_l_Lean_PPContext_runCoreM___rarg___closed__2(); @@ -3539,6 +4916,9 @@ l_Lean_PrettyPrinter_ppExprLegacy___closed__2 = _init_l_Lean_PrettyPrinter_ppExp lean_mark_persistent(l_Lean_PrettyPrinter_ppExprLegacy___closed__2); l_Lean_PrettyPrinter_ppExprLegacy___closed__3 = _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__3(); lean_mark_persistent(l_Lean_PrettyPrinter_ppExprLegacy___closed__3); +l_Lean_PrettyPrinter_ppExprLegacy___closed__4 = _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_ppExprLegacy___closed__4); +l_Lean_PrettyPrinter_ppExprLegacy___closed__5 = _init_l_Lean_PrettyPrinter_ppExprLegacy___closed__5(); l_Lean_PrettyPrinter_ppTactic___closed__1 = _init_l_Lean_PrettyPrinter_ppTactic___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_ppTactic___closed__1); l_Lean_PrettyPrinter_ppTactic___closed__2 = _init_l_Lean_PrettyPrinter_ppTactic___closed__2(); @@ -3553,52 +4933,52 @@ l_Lean_PrettyPrinter_ppModule___closed__2 = _init_l_Lean_PrettyPrinter_ppModule_ lean_mark_persistent(l_Lean_PrettyPrinter_ppModule___closed__2); l_Lean_PrettyPrinter_ppSignature___closed__1 = _init_l_Lean_PrettyPrinter_ppSignature___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_ppSignature___closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____lambda__1___closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__3); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__4); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__5); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739____closed__6); -if (builtin) {res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_739_(lean_io_mk_world()); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____lambda__1___closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__2); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__4); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__5); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746____closed__6); +if (builtin) {res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_746_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__1); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__2); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__3); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__4); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__5); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__6); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__7(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__7); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__8(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__8); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__9(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__9); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__10(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__10); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__11(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__11); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__12(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__12); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__13 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__13(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__13); -l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__14 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__14(); -lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807____closed__14); -if (builtin) {res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_807_(lean_io_mk_world()); +}l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__1); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__2 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__2); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__3); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__4 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__4); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__5 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__5); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__6 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__6); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__7 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__7(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__7); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__8 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__8(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__8); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__9 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__9(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__9); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__10 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__10(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__10); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__11 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__11(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__11); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__12 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__12(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__12); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__13 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__13(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__13); +l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__14 = _init_l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__14(); +lean_mark_persistent(l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814____closed__14); +if (builtin) {res = l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter___hyg_814_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_PrettyPrinter_registerParserCompilers___closed__1 = _init_l_Lean_PrettyPrinter_registerParserCompilers___closed__1(); @@ -3624,8 +5004,8 @@ lean_mark_persistent(l_Lean_PrettyPrinter_registerParserCompilers___closed__10); if (builtin) {res = l_Lean_PrettyPrinter_registerParserCompilers(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1 = _init_l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_MessageData_ofFormatWithInfos___lambda__1___closed__1); +}l_Lean_MessageData_ofFormatWithInfos___lambda__2___closed__1 = _init_l_Lean_MessageData_ofFormatWithInfos___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_MessageData_ofFormatWithInfos___lambda__2___closed__1); l_Lean_MessageData_ofFormatWithInfos___closed__1 = _init_l_Lean_MessageData_ofFormatWithInfos___closed__1(); lean_mark_persistent(l_Lean_MessageData_ofFormatWithInfos___closed__1); l_Lean_MessageData_ofConst___lambda__1___closed__1 = _init_l_Lean_MessageData_ofConst___lambda__1___closed__1(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c index 09c8c9b82ea2..c5e10dffc642 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Basic.c @@ -51,6 +51,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_addFieldInfo___boxed(lean_object*, lean_object*, lean_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_mkDelabAttribute___closed__13; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getPPOption(lean_object*); +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_init___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); @@ -65,10 +66,12 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_toStrin static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___closed__8; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAnonymous(lean_object*); lean_object* l_Lean_instMonadQuotationOfMonadFunctorOfMonadLift___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_shouldOmitExpr___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__1___closed__2; +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1; static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadReaderOfSubExprDelabM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,6 +91,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_noConfu LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_3938____closed__3; lean_object* l_Lean_getPPAnalysisNeedsType___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_toString___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_omission___closed__4; @@ -104,6 +108,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_D LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withIncDepth___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withOptionAtCurrPos___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_delab___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_OmissionReason_toString___boxed(lean_object*); lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___closed__14; @@ -111,16 +116,21 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_State_holeIter___defau uint8_t l_Lean_LocalContext_usesUserName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___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_PrettyPrinter_Delaborator_getExprKind___closed__18; +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delab___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isAtomic(lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadWithReaderOfSubExprDelabM___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_instAlternativeDelabM___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadStateOfHoleIteratorDelabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_SubExpr_Pos_maxChildren; static lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName___closed__2; +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3; +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_PrettyPrinter_Delaborator_addTermInfo___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_shouldOmitProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName___closed__3; @@ -156,6 +166,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_L LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName_bodyUsesSuggestion___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delab___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_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_getUnusedName_bodyUsesSuggestion___lambda__1(lean_object*, lean_object*, lean_object*); @@ -229,11 +240,14 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_L static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_shouldOmitProof___lambda__3___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__16; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_mkAppUnexpanderAttribute(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___closed__2; lean_object* l_Lean_RBNode_find___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_diagnostics; static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_addOmissionInfo___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_omission___closed__2; @@ -242,7 +256,7 @@ lean_object* lean_mk_syntax_ident(lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_HoleIterator_next(lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___closed__5; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___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_PrettyPrinter_delabCore___rarg___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_PrettyPrinter_delabCore___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_Option_get_x3f___at_Lean_PrettyPrinter_delabCore___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__1___closed__3; @@ -310,7 +324,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__20; static lean_object* l_Lean_PrettyPrinter_Delaborator_instAlternativeDelabM___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_mkDelabAttribute___closed__12; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instMonadReaderOfSubExprDelabM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___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_delabCore___rarg___lambda__4(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_PrettyPrinter_Delaborator_instMonadStateOfHoleIteratorDelabM___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_PrettyPrinter_Delaborator_shouldOmitExpr___lambda__2___closed__1; lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); @@ -335,6 +349,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instAlternativeDelabM; lean_object* l_Lean_Meta_eraseInaccessibleAnnotations___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_getExprKind___closed__2; LEAN_EXPORT 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_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_failure(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delab___lambda__2___closed__9; static lean_object* l_Lean_PrettyPrinter_initFn____x40_Lean_PrettyPrinter_Delaborator_Basic___hyg_3938____closed__11; @@ -11034,56 +11049,65 @@ return x_21; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___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) { +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1() { _start: { -uint8_t x_11; uint8_t x_12; -lean_dec(x_5); -x_11 = l_Lean_getPPInstantiateMVars(x_4); -x_12 = !lean_is_exclusive(x_8); +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(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, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; +x_12 = !lean_is_exclusive(x_9); if (x_12 == 0) { -lean_object* x_13; -x_13 = lean_ctor_get(x_8, 2); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 4); lean_dec(x_13); -lean_inc(x_4); -lean_ctor_set(x_8, 2, x_4); -if (x_11 == 0) -{ -lean_object* x_14; -x_14 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__3(x_4, x_1, x_2, x_3, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_4); -return x_14; +x_14 = lean_ctor_get(x_9, 2); +lean_dec(x_14); +x_15 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_15); +lean_ctor_set(x_9, 4, x_16); +lean_ctor_set(x_9, 2, x_1); +lean_ctor_set_uint8(x_9, sizeof(void*)*11 + 1, x_2); +if (x_3 == 0) +{ +lean_object* x_17; +x_17 = lean_apply_6(x_4, x_5, x_6, x_7, x_9, x_10, x_11); +return x_17; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_15 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_3, x_6, x_7, x_8, x_9, x_10); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__3(x_4, x_1, x_2, x_16, x_6, x_7, x_8, x_9, x_17); -lean_dec(x_4); -return x_18; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_5, x_6, x_7, 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_apply_6(x_4, x_19, x_6, x_7, x_9, x_10, x_20); +return x_21; } } 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; uint8_t x_29; uint8_t x_30; lean_object* x_31; -x_19 = lean_ctor_get(x_8, 0); -x_20 = lean_ctor_get(x_8, 1); -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); -x_30 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_22 = lean_ctor_get(x_9, 0); +x_23 = lean_ctor_get(x_9, 1); +x_24 = lean_ctor_get(x_9, 3); +x_25 = lean_ctor_get(x_9, 5); +x_26 = lean_ctor_get(x_9, 6); +x_27 = lean_ctor_get(x_9, 7); +x_28 = lean_ctor_get(x_9, 8); +x_29 = lean_ctor_get(x_9, 9); +x_30 = lean_ctor_get(x_9, 10); +x_31 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); +lean_inc(x_30); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_inc(x_26); @@ -11091,49 +11115,254 @@ lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_8); +lean_dec(x_9); +x_32 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1; +x_33 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_32); +x_34 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_34, 0, x_22); +lean_ctor_set(x_34, 1, x_23); +lean_ctor_set(x_34, 2, x_1); +lean_ctor_set(x_34, 3, x_24); +lean_ctor_set(x_34, 4, x_33); +lean_ctor_set(x_34, 5, x_25); +lean_ctor_set(x_34, 6, x_26); +lean_ctor_set(x_34, 7, x_27); +lean_ctor_set(x_34, 8, x_28); +lean_ctor_set(x_34, 9, x_29); +lean_ctor_set(x_34, 10, x_30); +lean_ctor_set_uint8(x_34, sizeof(void*)*11, x_31); +lean_ctor_set_uint8(x_34, sizeof(void*)*11 + 1, x_2); +if (x_3 == 0) +{ +lean_object* x_35; +x_35 = lean_apply_6(x_4, x_5, x_6, x_7, x_34, x_10, x_11); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_5, x_6, x_7, x_34, x_10, x_11); +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_apply_6(x_4, x_37, x_6, x_7, x_34, x_10, x_38); +return x_39; +} +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3; +x_2 = lean_unsigned_to_nat(0u); +x_3 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; uint8_t x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; +lean_dec(x_5); lean_inc(x_4); -x_31 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -lean_ctor_set(x_31, 2, x_4); -lean_ctor_set(x_31, 3, x_21); -lean_ctor_set(x_31, 4, x_22); -lean_ctor_set(x_31, 5, x_23); -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); -lean_ctor_set_uint8(x_31, sizeof(void*)*11 + 1, x_30); -if (x_11 == 0) +x_11 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___boxed), 9, 3); +lean_closure_set(x_11, 0, x_4); +lean_closure_set(x_11, 1, x_1); +lean_closure_set(x_11, 2, x_2); +x_12 = l_Lean_getPPInstantiateMVars(x_4); +x_13 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1; +x_14 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_4, x_13); +x_15 = lean_st_ref_get(x_9, x_10); +if (x_12 == 0) { -lean_object* x_32; -x_32 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__3(x_4, x_1, x_2, x_3, x_6, x_7, x_31, x_9, x_10); -lean_dec(x_4); -return x_32; +lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_55 = lean_ctor_get(x_15, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_15, 1); +lean_inc(x_56); +lean_dec(x_15); +x_57 = 0; +x_16 = x_57; +x_17 = x_55; +x_18 = x_56; +goto block_54; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(x_3, x_6, x_7, x_31, x_9, x_10); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_ctor_get(x_15, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_15, 1); +lean_inc(x_59); +lean_dec(x_15); +x_60 = 1; +x_16 = x_60; +x_17 = x_58; +x_18 = x_59; +goto block_54; +} +block_54: +{ +lean_object* x_19; uint8_t x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Kernel_isDiagnosticsEnabled(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +if (x_14 == 0) +{ +uint8_t x_50; +x_50 = 1; +x_21 = x_50; +goto block_49; +} +else +{ +uint8_t x_51; +x_51 = 0; +x_21 = x_51; +goto block_49; +} +} +else +{ +if (x_14 == 0) +{ +uint8_t x_52; +x_52 = 0; +x_21 = x_52; +goto block_49; +} +else +{ +uint8_t x_53; +x_53 = 1; +x_21 = x_53; +goto block_49; +} +} +block_49: +{ +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_st_ref_take(x_9, x_18); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_26 = lean_ctor_get(x_23, 0); +x_27 = lean_ctor_get(x_23, 4); +lean_dec(x_27); +x_28 = l_Lean_Kernel_enableDiag(x_26, x_14); +x_29 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5; +lean_ctor_set(x_23, 4, x_29); +lean_ctor_set(x_23, 0, x_28); +x_30 = lean_st_ref_set(x_9, x_23, x_24); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_box(0); +x_33 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_4, x_14, x_16, x_11, x_3, x_6, x_7, x_32, x_8, x_9, 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; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_34 = lean_ctor_get(x_23, 0); +x_35 = lean_ctor_get(x_23, 1); +x_36 = lean_ctor_get(x_23, 2); +x_37 = lean_ctor_get(x_23, 3); +x_38 = lean_ctor_get(x_23, 5); +x_39 = lean_ctor_get(x_23, 6); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__3(x_4, x_1, x_2, x_34, x_6, x_7, x_31, x_9, x_35); -lean_dec(x_4); -return x_36; +lean_inc(x_34); +lean_dec(x_23); +x_40 = l_Lean_Kernel_enableDiag(x_34, x_14); +x_41 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5; +x_42 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_35); +lean_ctor_set(x_42, 2, x_36); +lean_ctor_set(x_42, 3, x_37); +lean_ctor_set(x_42, 4, x_41); +lean_ctor_set(x_42, 5, x_38); +lean_ctor_set(x_42, 6, x_39); +x_43 = lean_st_ref_set(x_9, x_42, x_24); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_box(0); +x_46 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_4, x_14, x_16, x_11, x_3, x_6, x_7, x_45, x_8, x_9, x_44); +return x_46; } } +else +{ +lean_object* x_47; lean_object* x_48; +x_47 = lean_box(0); +x_48 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_4, x_14, x_16, x_11, x_3, x_6, x_7, x_47, x_8, x_9, x_18); +return x_48; } } -static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1() { +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1() { _start: { lean_object* x_1; @@ -11141,14 +11370,14 @@ x_1 = l_Lean_pp_proofs; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___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) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_dec(x_4); x_10 = lean_ctor_get(x_7, 2); lean_inc(x_10); -x_11 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1; +x_11 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1; x_12 = l_Lean_Option_get_x3f___at_Lean_PrettyPrinter_delabCore___spec__3(x_10, x_11); x_13 = lean_box(0); x_14 = l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_PrettyPrinter_delabCore___spec__4(x_12, x_13); @@ -11157,7 +11386,7 @@ if (x_14 == 0) { lean_object* x_15; lean_object* x_16; x_15 = lean_box(0); -x_16 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_15, x_5, x_6, x_7, x_8, x_9); +x_16 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_15, x_5, x_6, x_7, x_8, x_9); return x_16; } else @@ -11187,7 +11416,7 @@ 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_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_22, x_5, x_6, x_7, x_8, x_21); +x_23 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_22, x_5, x_6, x_7, x_8, x_21); return x_23; } else @@ -11199,7 +11428,7 @@ lean_dec(x_18); x_25 = 1; x_26 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_10, x_11, x_25); x_27 = lean_box(0); -x_28 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_26, x_27, x_5, x_6, x_7, x_8, x_24); +x_28 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_26, x_27, x_5, x_6, x_7, x_8, x_24); return x_28; } } @@ -11219,7 +11448,7 @@ lean_object* x_33; lean_object* x_34; lean_free_object(x_18); lean_dec(x_30); x_33 = lean_box(0); -x_34 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_33, x_5, x_6, x_7, x_8, x_31); +x_34 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_33, x_5, x_6, x_7, x_8, x_31); return x_34; } else @@ -11244,7 +11473,7 @@ lean_object* x_36; lean_object* x_37; lean_free_object(x_18); lean_dec(x_30); x_36 = lean_box(0); -x_37 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_36, x_5, x_6, x_7, x_8, x_31); +x_37 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_36, x_5, x_6, x_7, x_8, x_31); return x_37; } } @@ -11263,7 +11492,7 @@ if (x_40 == 0) lean_object* x_41; lean_object* x_42; lean_dec(x_38); x_41 = lean_box(0); -x_42 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_41, x_5, x_6, x_7, x_8, x_39); +x_42 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_41, x_5, x_6, x_7, x_8, x_39); return x_42; } else @@ -11291,7 +11520,7 @@ else lean_object* x_45; lean_object* x_46; lean_dec(x_38); x_45 = lean_box(0); -x_46 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_45, x_5, x_6, x_7, x_8, x_39); +x_46 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_45, x_5, x_6, x_7, x_8, x_39); return x_46; } } @@ -11302,7 +11531,7 @@ else { lean_object* x_47; lean_object* x_48; x_47 = lean_box(0); -x_48 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_2, x_3, x_10, x_47, x_5, x_6, x_7, x_8, x_9); +x_48 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_1, x_2, x_3, x_10, x_47, x_5, x_6, x_7, x_8, x_9); return x_48; } } @@ -11390,7 +11619,7 @@ 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_PrettyPrinter_delabCore___rarg___lambda__5(x_3, x_2, x_12, x_19, x_4, x_5, x_6, x_7, x_18); +x_20 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__6(x_3, x_2, x_12, x_19, x_4, x_5, x_6, x_7, x_18); return x_20; } else @@ -11417,7 +11646,7 @@ lean_inc(x_29); x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__5(x_3, x_2, x_12, x_29, x_4, x_5, x_6, x_7, x_30); +x_31 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__6(x_3, x_2, x_12, x_29, x_4, x_5, x_6, x_7, x_30); return x_31; } } @@ -11507,6 +11736,19 @@ lean_dec(x_1); return x_10; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_12 = lean_unbox(x_2); +lean_dec(x_2); +x_13 = lean_unbox(x_3); +lean_dec(x_3); +x_14 = l_Lean_PrettyPrinter_delabCore___rarg___lambda__4(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_8); +return x_14; +} +} LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_delab(lean_object* x_1, 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: { @@ -11996,8 +12238,20 @@ l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___closed__1 = _init_l_Lean_Pre lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___closed__1); l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__3___closed__2); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__4___closed__1); l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__1); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__2); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__3); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__4); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__5___closed__5); +l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1 = _init_l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___lambda__6___closed__1); l_Lean_PrettyPrinter_delabCore___rarg___closed__1 = _init_l_Lean_PrettyPrinter_delabCore___rarg___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_delabCore___rarg___closed__1); l_Lean_PrettyPrinter_delabCore___rarg___closed__2 = _init_l_Lean_PrettyPrinter_delabCore___rarg___closed__2(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c index 4feb115ceac3..4c846613b097 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/Builtins.c @@ -14,17 +14,16 @@ extern "C" { #endif 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_delabLetFun___lambda__2___closed__5; 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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__5; lean_object* l_Lean_Expr_bindingName_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_params___default; LEAN_EXPORT lean_object* l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__1___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabForall___closed__5; lean_object* l_Lean_getPPMVars___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__6; static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1___closed__1; @@ -38,7 +37,6 @@ lean_object* lean_string_utf8_extract(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__11(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext_getPPBinderTypes(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getCoeFnInfo_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPStructureInstanceType___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__2; @@ -46,15 +44,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda 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*); +static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, 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*); 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_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_object*); -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___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_delabDIte(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx(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_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_declSigWithId___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__9___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__4; @@ -62,8 +60,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfNat(lean_object*); -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_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*); @@ -71,15 +67,15 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora 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_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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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_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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_nextExtraPos___at_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_needsExplicit___spec__1(lean_object*, size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_delabCond___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___closed__1; @@ -89,21 +85,25 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___lamb 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_PrettyPrinter_Delaborator_ParamKind_isRegularExplicit___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator(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_withoutParentProjections___spec__1___rarg(lean_object*, 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_Lean_PrettyPrinter_Delaborator_AppMatchState_alts___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__5; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__1; lean_object* l_Lean_Parser_mkAntiquot_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_delabAppImplicitCore___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_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch(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__9(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; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___closed__4; @@ -121,21 +121,21 @@ 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_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__7(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; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__15___closed__2; 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 uint8_t l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(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_instInhabitedParamKind; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__8___closed__1; 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*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3; -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_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_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__6___boxed(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; @@ -156,11 +156,9 @@ 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_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__3(lean_object*); 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; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1(lean_object*); +static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1; 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*); @@ -170,20 +168,26 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lamb 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*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___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_PrettyPrinter_Delaborator_delabAppImplicitCore___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___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_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte(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_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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13; +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*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1___boxed(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; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppCore___lambda__1(uint8_t, 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_delabLetFun___lambda__2___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__2; 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*); @@ -203,7 +207,7 @@ 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_delabAppCore___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_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___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_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*); @@ -211,16 +215,15 @@ lean_object* l_Lean_Syntax_getId(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; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(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_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_Array_findIdx_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNatCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); -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__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_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*); @@ -229,10 +232,14 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delabora 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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5___boxed(lean_object*, lean_object*, lean_object*); 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; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__7(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__1; +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___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_delabApp___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabChar___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -240,6 +247,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1_ 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_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_delabAppImplicitCore_mkArg___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_delabDivRatCore___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_withMDatasOptions(lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); @@ -253,27 +262,29 @@ uint8_t lean_usize_dec_eq(size_t, size_t); uint8_t l_Lean_getPPMotivesAll(lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_needsExplicit(lean_object*, lean_object*, lean_object*); 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*); +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_object*, lean_object*); uint8_t l_Lean_getPPFunBinderTypes(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArgs(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__4; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___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_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*); +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__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_SubExpr_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun(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_delabLetFun___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_PrettyPrinter_Delaborator_delabConst___closed__13; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__8___closed__2; 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_delabAppExplicitCore___lambda__6___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__2___closed__2; lean_object* lean_array_fget(lean_object*, lean_object*); @@ -285,32 +296,38 @@ 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_object* l_Lean_Meta_Match_MatcherInfo_arity(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_Lean_PrettyPrinter_Delaborator_AppMatchState_hNames_x3f___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_useAppExplicit___lambda__5___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__13(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_TSepArray_getElems___rarg(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default; uint8_t l_Lean_Expr_isLetFun(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_delabParamsAux___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___boxed(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; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___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_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__1; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__4; 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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___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_unexpandStructureInstance___lambda__6___closed__2; lean_object* lean_environment_find(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12; +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__8(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_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; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__5; +static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getBoundedAppFn(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__2___rarg___closed__2; @@ -318,12 +335,15 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId___closed__4; extern lean_object* l_Lean_ppUsingAnonymousConstructorAttr; 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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabOfScientific(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___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_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_appFieldNotationCandidate_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; 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_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*); @@ -332,7 +352,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withMDataOptions___rar 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*); @@ -340,11 +359,8 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_del LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___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_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___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__2(lean_object*, size_t, size_t, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2; 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*); @@ -363,17 +379,19 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x 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_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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_withOptionAtCurrPos___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_delabSort___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_delabSort___closed__9; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___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_delabForall___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1; +extern lean_object* l_Lean_maxRecDepth; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__2; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___closed__3; lean_object* l_Lean_KVMap_insertCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__6___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -388,73 +406,79 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort 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; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__1; lean_object* l_Lean_getPPUniverses___boxed(lean_object*); uint8_t l_Lean_Name_isPrefixOf(lean_object*, lean_object*); +static lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2; 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_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__2; -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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4___boxed(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_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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__1; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_evalSyntaxConstantUnsafe___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__8___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__6(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1; 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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_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*); +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__8; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders(lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__5; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___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_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedNat; -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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(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_getNumArgs(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar___closed__5; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPMatch___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getStructureFields(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___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_delabAppImplicitCore___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__7(lean_object*); 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*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__6; -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_delabAppImplicitCore_mkNamedArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_mkApp(lean_object*, lean_object*); LEAN_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_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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_unexpandStructureInstance___lambda__7___boxed(lean_object**); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNeg___closed__1; lean_object* l_Array_mapMUnsafe_map___at___aux__Init__NotationExtra______macroRules__term_x25_x5b___x7c___x5d__1___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___spec__17(lean_object*, lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7(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___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg___closed__4; lean_object* l_Lean_isLHSGoal_x3f(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabChar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPStructureInstancesFlatten___boxed(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__3; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___closed__4; -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*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; uint8_t l_Lean_Expr_isLambda(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; -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg(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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_getParamKinds___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___regBuiltin_Lean_PrettyPrinter_Delaborator_coeDelaborator___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___closed__2; @@ -465,33 +489,34 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv___cl 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*); 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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__2; 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; 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(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__6___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2___closed__4; 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; 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_appFieldNotationCandidate_x3f___lambda__2___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_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*); extern lean_object* l_Lean_Parser_Command_declSig; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_getParamKinds___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -506,6 +531,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___clos LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabProj___spec__12(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__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern(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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__3; @@ -514,6 +540,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_d 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_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23(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; @@ -523,12 +550,11 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__9___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___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_PrettyPrinter_Delaborator_withTypeAscription___closed__11; lean_object* l_Lean_getPPAnalysisLetVarType___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNeg___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__3; lean_object* l_Array_zip___rarg(lean_object*, lean_object*); @@ -541,11 +567,11 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___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_delabAppExplicitCore___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2; lean_object* l_Lean_Level_quote(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPStructureInstances___boxed(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__6___closed__4; uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -555,14 +581,18 @@ lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_ 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_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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(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_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_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__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_delabOfScientific___closed__1; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__16(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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_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_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__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_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; @@ -574,6 +604,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFora 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__1; 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*); @@ -585,12 +616,11 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1_ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withOverApp___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_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_PrettyPrinter_Delaborator_delabAppMatch_usingNames(lean_object*); 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*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___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*); 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*); 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*); @@ -603,10 +633,10 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___c 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_delabLetFun___lambda__2___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__3; 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*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5; LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___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_PrettyPrinter_Delaborator_delabCond(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__6; @@ -615,20 +645,24 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___clo LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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_delabSort___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__6; +lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(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_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___boxed(lean_object**); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_AppMatchState_motiveNamed___default; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabProj(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_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections___spec__1(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*); +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__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_delabLam___lambda__3___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabBinders___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___lambda__1___closed__2; @@ -636,18 +670,20 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__5; LEAN_EXPORT uint8_t l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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_PrettyPrinter_Delaborator_delabSigmaCore___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__7___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___closed__5; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9(lean_object*, size_t, size_t); lean_object* l_Lean_Name_replacePrefix(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_SubExpr_Pos_pushNthBindingDomain(lean_object*, lean_object*); 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; 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*); +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*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___lambda__1___closed__1; lean_object* lean_array_pop(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__3; @@ -658,18 +694,18 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambd static lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2___closed__1; 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*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabDoElems___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_leadingNode_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_pp_universes; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; lean_object* l_Lean_PrettyPrinter_Parenthesizer_andthen_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___closed__3; uint8_t l_List_isEmpty___rarg(lean_object*); 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___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_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_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_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_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*); @@ -679,7 +715,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__3; 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*); @@ -695,20 +730,19 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__14; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8(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_ParamKind_defVal___default; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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*); 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_annotateTermInfo___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_delabDIte___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_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3___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_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; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___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_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__5(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_maybeAddBlockImplicit___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabBVar___closed__1; @@ -717,10 +751,11 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__4; lean_object* l_Lean_PrettyPrinter_Delaborator_getUnusedName(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_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(size_t, size_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_withAnnotateTermInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -728,6 +763,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__8; 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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__2___boxed(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*); @@ -742,22 +778,27 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatte uint8_t l_Lean_BinderInfo_isInstImplicit(uint8_t); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2(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_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___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_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(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__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; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___lambda__7___closed__3; -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__4; 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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9___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_Name_componentsRev(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__8___closed__3; -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*); lean_object* l_Lean_getPPAnalysisNoDot___boxed(lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -768,6 +809,8 @@ lean_object* l_Lean_Name_str___override(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*); lean_object* l_Lean_getPPCoercions___boxed(lean_object*); extern lean_object* l_Lean_rootNamespace; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___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*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_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*); @@ -780,10 +823,11 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance__ static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__1; lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__6___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr(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_PrettyPrinter_Delaborator_coeDelaborator___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__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_getParamKinds(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_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; @@ -793,6 +837,7 @@ lean_object* l_Lean_getPPAnalysisHole___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__11; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -805,23 +850,27 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLam___clo 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_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___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*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2; 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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4(lean_object*, lean_object*, size_t, size_t, 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*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(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___lambda__2(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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabProj(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__6; +extern lean_object* l_Lean_diagnostics; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems(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_SubExpr_Pos_pushNaryArg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNeg___closed__3; -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections(lean_object*, lean_object*, lean_object*, 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; @@ -830,15 +879,17 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatur 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2___closed__6; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__3; +LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5(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__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___closed__2; extern lean_object* l_Lean_Meta_instMonadMetaM; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabDoElems___spec__3(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___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* l_Array_eraseIdx___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___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_PrettyPrinter_Delaborator_delabNamedPattern___closed__1; @@ -846,16 +897,19 @@ 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 uint8_t l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand(lean_object*); lean_object* l_Lean_getPPFieldNotationGeneralized___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__6___closed__3; 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; +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_getOptParamDefault_x3f(lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_ParamKind_isRegularExplicit(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMVar___closed__2; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__5; @@ -870,7 +924,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_ 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*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__3___closed__11; -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2; @@ -879,12 +932,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at__ lean_object* l_Lean_MessageData_ofExpr(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkNum(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__3; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6; 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_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*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__17(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__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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__3; 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*); @@ -896,6 +952,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore__ 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__3(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabHDiv___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -906,26 +963,25 @@ lean_object* l_Lean_Syntax_mkCharLit(uint32_t, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; static lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; uint8_t l_Lean_getPPMotivesNonConst(lean_object*); -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__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__3(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_PrettyPrinter_Delaborator_shouldShowMotive___closed__1; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; lean_object* l_Lean_LocalDecl_userName(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; lean_object* l_Lean_getPPMVarsWithType___boxed(lean_object*); 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*); uint8_t l_Lean_Expr_isAutoParam(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1; 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*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__5; 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*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabHDiv(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_getPPPiBinderTypes(lean_object*); @@ -934,7 +990,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_shouldShowMotive___box 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; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___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_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__1; lean_object* l_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent(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__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -951,14 +1006,16 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore__ 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_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___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_object*); uint8_t l_Lean_Name_hasMacroScopes(lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___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_delabAppImplicitCore___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_Parser_leadingNode_formatter___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_withOverApp___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_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___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_delabFVar___closed__7; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__15(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNameMkStr(lean_object*); lean_object* l_Lean_PrettyPrinter_Parenthesizer_withAntiquot_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -970,13 +1027,14 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1( static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabDo___spec__1(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__5; 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; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__3; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2; lean_object* l_Lean_Name_appendCore(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__6; @@ -995,8 +1053,10 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_n LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___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___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__4; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_returnsPi___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_hasIdent(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; @@ -1011,13 +1071,15 @@ lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lea static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppCore___lambda__1___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond___closed__1; lean_object* l_Lean_getPPPiBinderTypes___boxed(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4; +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_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*); 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; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___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_delabAppExplicitCore___lambda__8___closed__3; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; @@ -1026,12 +1088,15 @@ 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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2; 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_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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_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_object* l_Lean_Parser_withCache(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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__1; 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*); @@ -1039,7 +1104,6 @@ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPSigma___ 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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___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_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5; @@ -1059,21 +1123,22 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__6; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabFVar(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__9; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__6; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__9(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__7(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___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_hasIdent___boxed(lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections(lean_object*); 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_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___closed__4; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2; 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*); @@ -1082,17 +1147,19 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance__ lean_object* l_Lean_Parser_Command_declSig_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__1(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetFun(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(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_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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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_delabAppMatch___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabCond___closed__2; -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_object* lean_array_set(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*); @@ -1101,7 +1168,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore__ 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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabHDiv___closed__2; lean_object* l_Lean_Level_dec(lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalDecl_type(lean_object*); +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabForall___spec__1___closed__2; uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); @@ -1112,6 +1181,7 @@ lean_object* l_Lean_getPPInstances___boxed(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___closed__1; 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_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_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___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_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___regBuiltin_Lean_PrettyPrinter_Delaborator_delabChar(lean_object*); @@ -1128,19 +1198,22 @@ lean_object* l_List_tailD___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__1___closed__8; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__4; 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_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___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_instInhabitedSubExpr; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___boxed(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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(lean_object*, size_t, size_t, 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_delabAppImplicitCore___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19(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_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNamedPattern(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1153,31 +1226,35 @@ LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__ static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__3; 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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__2___boxed(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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; 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_delabDoElems___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; 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*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLit___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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_object*); 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*); 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*); +static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__1___closed__2; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__4; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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_delabLit___closed__3; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSort___closed__2; lean_object* l_Array_ofSubarray___rarg(lean_object*); 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*); @@ -1185,44 +1262,42 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__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_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__2(lean_object*, size_t, size_t, lean_object*); 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; static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__4; 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_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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__2; lean_object* l_Array_back___rarg(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_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(lean_object*); 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_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___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_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_delabAppImplicitCore_mkArgStx___closed__6; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(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; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__7___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_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__4(lean_object*); size_t lean_usize_sub(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_delabAppMatch___lambda__3___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNegIntCore___closed__7; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabBVar___closed__1; -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__1___closed__3; 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; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1; 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*); @@ -1235,24 +1310,21 @@ 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*); +static lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1; 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_delabAppImplicitCore___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_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(size_t, size_t, lean_object*); +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3(lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabApp___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node8(lean_object*, 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_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___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_skippingBinders_loop_visitLambda___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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__2; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2(lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNames___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_delabLam___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_PrettyPrinter_Delaborator_withTypeAscription___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; @@ -1262,11 +1334,14 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___clo 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*); lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___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*, 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_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(uint8_t, lean_object*, 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; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__10(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_reifyName___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1277,20 +1352,22 @@ lean_object* l_Lean_PrettyPrinter_Delaborator_addFieldInfo(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabLam___spec__3(lean_object*, size_t, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__5; extern lean_object* l_Lean_PrettyPrinter_formatterAttribute; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLam___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_redLength___rarg(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPExplicit___boxed(lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo___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_coeDelaborator___lambda__4___boxed(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*, lean_object*, 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_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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1; 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*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___lambda__1___closed__2; @@ -1298,14 +1375,15 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatur 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*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1; 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_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_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__6(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2; 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*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20___boxed(lean_object*, lean_object*, lean_object*); 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*); @@ -1313,29 +1391,28 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_delabSort___closed__1; 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*); 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*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22___boxed(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_withoutParentProjections___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t 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; +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; 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; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3; 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*); lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabSigma___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10(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_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3; +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_Lean_PrettyPrinter_Delaborator_declSigWithId_parenthesizer___closed__2; 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*); +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*); uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1; lean_object* l_Array_mkArray2___rarg(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfNat___closed__5; @@ -1348,12 +1425,10 @@ lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint lean_object* l_Lean_LocalContext_getUnusedName(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___lambda__2___closed__3; extern lean_object* l_Lean_PrettyPrinter_Delaborator_appUnexpanderAttribute; -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___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_unresolveNameGlobal___at_Lean_PrettyPrinter_Delaborator_delabConst___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_PrettyPrinter_Delaborator_delabConst___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_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabNamedPattern___closed__3; @@ -1365,6 +1440,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___c uint8_t l_Lean_Expr_isConst(lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabOfScientific___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_withOverApp___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1374,8 +1450,6 @@ lean_object* l_Lean_Syntax_mkScientificLit(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__9___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabConst___closed__10; static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_withMDataOptions___spec__3___rarg___closed__1; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; static lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__6___closed__4; static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_declSigWithId_formatter___closed__1; @@ -1387,16 +1461,18 @@ 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_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_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(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_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; 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*); +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*, lean_object*); lean_object* l_Lean_Expr_consumeMData(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabForall___lambda__1___closed__2; -static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1; +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabBVar(lean_object*); +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__5; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetE___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_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1; @@ -1405,6 +1481,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabLetE___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getPPAnalysisBlockImplicit___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(size_t, size_t, lean_object*); 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*); @@ -1417,14 +1494,14 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withProj___at_Lean_ 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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___spec__1(lean_object*, size_t, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_delabAppMatch___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___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__6; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDo(lean_object*); -static lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__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_unexpandStructureInstance___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__5; lean_object* l_Char_ofNat(lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4; @@ -1453,12 +1530,13 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabDIte___closed__5; static lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1___closed__4; 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*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; lean_object* l_Lean_Parser_termParser(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__2; uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__3; static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1() { @@ -11007,6 +11085,294 @@ x_11 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter return x_11; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appArg_x21(x_10); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; +} +} +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("proj", 4); +return x_1; +} +} +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ident", 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections(lean_object* x_1, lean_object* x_2, 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 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; +lean_inc(x_1); +x_10 = l_Lean_Syntax_isOfKind(x_1, x_9); +if (x_10 == 0) +{ +lean_object* x_11; +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_8); +return x_11; +} +else +{ +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(0u); +x_13 = l_Lean_Syntax_getArg(x_1, x_12); +x_14 = lean_unsigned_to_nat(2u); +x_15 = l_Lean_Syntax_getArg(x_1, x_14); +x_16 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; +lean_inc(x_15); +x_17 = l_Lean_Syntax_isOfKind(x_15, x_16); +if (x_17 == 0) +{ +lean_object* x_18; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_1); +lean_ctor_set(x_18, 1, x_8); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; lean_object* x_33; +x_29 = 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_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 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_33 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(x_32, x_30, x_4, x_5, x_6, x_7, x_31); +if (lean_obj_tag(x_33) == 0) +{ +lean_object* x_34; lean_object* x_35; +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_19 = x_34; +x_20 = x_35; +goto block_28; +} +else +{ +uint8_t x_36; +x_36 = !lean_is_exclusive(x_33); +if (x_36 == 0) +{ +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_ctor_get(x_33, 0); +x_38 = lean_ctor_get(x_33, 1); +x_39 = l_Lean_Exception_isRuntime(x_37); +if (x_39 == 0) +{ +lean_object* x_40; +lean_free_object(x_33); +lean_dec(x_37); +x_40 = lean_box(0); +x_19 = x_40; +x_20 = x_38; +goto block_28; +} +else +{ +uint8_t x_41; +x_41 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +if (x_41 == 0) +{ +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_33; +} +else +{ +lean_object* x_42; +lean_free_object(x_33); +lean_dec(x_37); +x_42 = lean_box(0); +x_19 = x_42; +x_20 = x_38; +goto block_28; +} +} +} +else +{ +lean_object* x_43; lean_object* x_44; uint8_t x_45; +x_43 = lean_ctor_get(x_33, 0); +x_44 = lean_ctor_get(x_33, 1); +lean_inc(x_44); +lean_inc(x_43); +lean_dec(x_33); +x_45 = l_Lean_Exception_isRuntime(x_43); +if (x_45 == 0) +{ +lean_object* x_46; +lean_dec(x_43); +x_46 = lean_box(0); +x_19 = x_46; +x_20 = x_44; +goto block_28; +} +else +{ +uint8_t x_47; +x_47 = lean_ctor_get_uint8(x_6, sizeof(void*)*11); +if (x_47 == 0) +{ +lean_object* x_48; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_48 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_48, 0, x_43); +lean_ctor_set(x_48, 1, x_44); +return x_48; +} +else +{ +lean_object* x_49; +lean_dec(x_43); +x_49 = lean_box(0); +x_19 = x_49; +x_20 = x_44; +goto block_28; +} +} +} +} +block_28: +{ +if (lean_obj_tag(x_19) == 0) +{ +lean_object* x_21; +lean_dec(x_15); +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_20); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_19, 0); +lean_inc(x_22); +lean_dec(x_19); +x_23 = l_Lean_Syntax_getId(x_15); +lean_dec(x_15); +x_24 = lean_name_eq(x_23, x_22); +lean_dec(x_22); +lean_dec(x_23); +if (x_24 == 0) +{ +lean_object* x_25; +lean_dec(x_13); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_20); +return x_25; +} +else +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_1); +x_26 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections), 8, 1); +lean_closure_set(x_26, 0, x_13); +x_27 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +return x_27; +} +} +} +} +} +} +} LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_needsExplicit___spec__1(lean_object* x_1, size_t x_2, size_t x_3) { _start: { @@ -12465,26 +12831,6 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore_ _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("proj", 4); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___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_delabAppExplicitCore___lambda__6___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_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes(".", 1); return x_1; } @@ -12595,13 +12941,13 @@ x_40 = lean_ctor_get(x_10, 5); lean_inc(x_40); x_41 = 0; x_42 = l_Lean_SourceInfo_fromRef(x_40, x_41); -x_43 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +x_43 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; lean_inc(x_42); x_44 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_44, 0, x_42); lean_ctor_set(x_44, 1, x_43); x_45 = lean_mk_syntax_ident(x_35); -x_46 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2; +x_46 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; x_47 = l_Lean_Syntax_node3(x_42, x_46, x_39, x_44, x_45); x_48 = lean_unsigned_to_nat(0u); x_49 = lean_nat_dec_eq(x_34, x_48); @@ -12692,13 +13038,13 @@ x_65 = lean_ctor_get(x_10, 5); lean_inc(x_65); x_66 = 0; x_67 = l_Lean_SourceInfo_fromRef(x_65, x_66); -x_68 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +x_68 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; lean_inc(x_67); x_69 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_69, 0, x_67); lean_ctor_set(x_69, 1, x_68); x_70 = lean_mk_syntax_ident(x_35); -x_71 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2; +x_71 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; x_72 = l_Lean_Syntax_node3(x_67, x_71, x_64, x_69, x_70); x_73 = lean_unsigned_to_nat(0u); x_74 = lean_nat_dec_eq(x_34, x_73); @@ -13426,6 +13772,70 @@ x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore(x_12, x_2, x_3, x_4 return x_13; } } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg() { +_start: +{ +lean_object* x_1; +x_1 = lean_box(0); +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 2) +{ +uint8_t x_2; +x_2 = 0; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 1; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(lean_object* x_1) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_2; +x_2 = lean_box(0); +return x_2; +} +else +{ +lean_object* x_3; lean_object* x_4; +x_3 = lean_ctor_get(x_1, 0); +lean_inc(x_3); +x_4 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_4, 0, x_3); +return x_4; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_1); +lean_dec(x_1); +return x_2; +} +} static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__1() { _start: { @@ -13446,98 +13856,122 @@ 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_delabAppImplicitCore_mkNamedArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_delabAppImplicitCore_mkNamedArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; 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; -x_10 = lean_ctor_get(x_7, 5); -lean_inc(x_10); -lean_dec(x_7); -x_11 = 0; -x_12 = l_Lean_SourceInfo_fromRef(x_10, x_11); -x_13 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; -lean_inc(x_12); -x_14 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_14, 0, x_12); -lean_ctor_set(x_14, 1, x_13); -x_15 = lean_mk_syntax_ident(x_1); -x_16 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; -lean_inc(x_12); -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_12); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_object* x_9; +lean_inc(x_6); +x_9 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_6, 5); lean_inc(x_12); +lean_dec(x_6); +x_13 = 0; +x_14 = l_Lean_SourceInfo_fromRef(x_12, x_13); +x_15 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +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 = lean_mk_syntax_ident(x_1); +x_18 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; +lean_inc(x_14); x_19 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_19, 0, x_12); +lean_ctor_set(x_19, 0, x_14); lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2; -x_21 = l_Lean_Syntax_node5(x_12, x_20, x_14, x_15, x_17, x_2, x_19); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_21); -x_23 = lean_box(x_11); -x_24 = lean_alloc_ctor(0, 2, 0); +x_20 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_14); +x_21 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_20); +x_22 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2; +x_23 = l_Lean_Syntax_node5(x_14, x_22, x_16, x_17, x_19, x_11, x_21); +x_24 = lean_alloc_ctor(2, 1, 0); lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, 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; -} +lean_ctor_set(x_9, 0, x_24); +return x_9; } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); +lean_object* x_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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_25 = lean_ctor_get(x_9, 0); +x_26 = lean_ctor_get(x_9, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_9); +x_27 = lean_ctor_get(x_6, 5); +lean_inc(x_27); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_10; +x_28 = 0; +x_29 = l_Lean_SourceInfo_fromRef(x_27, x_28); +x_30 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +lean_inc(x_29); +x_31 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_31, 0, x_29); +lean_ctor_set(x_31, 1, x_30); +x_32 = lean_mk_syntax_ident(x_1); +x_33 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; +lean_inc(x_29); +x_34 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_34, 0, x_29); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_29); +x_36 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_36, 0, x_29); +lean_ctor_set(x_36, 1, x_35); +x_37 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2; +x_38 = l_Lean_Syntax_node5(x_29, x_37, x_31, x_32, x_34, x_25, x_36); +x_39 = lean_alloc_ctor(2, 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_26); +return x_40; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1___closed__1; -x_10 = lean_panic_fn(x_9, x_1); -x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_11; -} +uint8_t x_41; +lean_dec(x_6); +lean_dec(x_1); +x_41 = !lean_is_exclusive(x_9); +if (x_41 == 0) +{ +return x_9; } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_getPPAnalysisSkip___boxed), 1, 0); -return x_1; +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_9, 0); +x_43 = lean_ctor_get(x_9, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_9); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Lean.PrettyPrinter.Delaborator.delabAppImplicitCore.mkArgStx", 60); -return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___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_PrettyPrinter_Delaborator_delabFVar___closed__1; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2; -x_3 = lean_unsigned_to_nat(391u); -x_4 = lean_unsigned_to_nat(37u); -x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_getPPAnalysisSkip___boxed), 1, 0); +return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2() { _start: { lean_object* x_1; @@ -13545,7 +13979,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_getPPAnalysisNamedArg___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3() { _start: { lean_object* x_1; @@ -13553,1773 +13987,1013 @@ x_1 = lean_mk_string_from_bytes("motive", 6); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; lean_object* x_4; -x_1 = lean_box(0); -x_2 = 1; -x_3 = lean_box(x_2); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_1); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_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_delabAppImplicitCore_mkArg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_11 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) +x_14 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_12; uint8_t x_13; -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_unbox(x_12); -lean_dec(x_12); -if (x_13 == 0) +lean_object* x_15; uint8_t x_16; +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_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); -lean_dec(x_11); -x_15 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__1; +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_14, 1); +lean_inc(x_17); +lean_dec(x_14); +x_18 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_16 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; uint8_t x_18; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_unbox(x_17); -lean_dec(x_17); -if (x_18 == 0) -{ -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 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_2) == 0) +x_19 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_1); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); +lean_object* x_20; uint8_t x_21; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_unbox(x_20); lean_dec(x_20); -x_22 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3; -x_23 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___spec__1(x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_21); -return x_23; -} -else +if (x_21 == 0) { -uint8_t x_24; -x_24 = !lean_is_exclusive(x_20); -if (x_24 == 0) +uint8_t x_22; +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_25 = lean_ctor_get(x_20, 0); -x_26 = lean_ctor_get(x_20, 1); -x_27 = lean_ctor_get(x_2, 0); -lean_inc(x_27); -lean_dec(x_2); -x_28 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; +x_23 = lean_ctor_get(x_19, 1); +x_24 = lean_ctor_get(x_19, 0); +lean_dec(x_24); +x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_29 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_26); -if (lean_obj_tag(x_29) == 0) +x_26 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_23); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_63; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_29)) { - lean_ctor_release(x_29, 0); - lean_ctor_release(x_29, 1); - x_32 = x_29; +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_67; +x_27 = lean_ctor_get(x_26, 0); +lean_inc(x_27); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_29 = x_26; } else { - lean_dec_ref(x_29); - x_32 = lean_box(0); -} -x_63 = lean_unbox(x_30); -lean_dec(x_30); -if (x_63 == 0) -{ -lean_object* x_64; -x_64 = lean_ctor_get(x_27, 1); -lean_inc(x_64); -if (lean_obj_tag(x_64) == 0) -{ -uint8_t x_65; uint8_t x_66; -lean_free_object(x_20); -lean_dec(x_1); -x_65 = lean_ctor_get_uint8(x_27, sizeof(void*)*2); -x_66 = l_Lean_BinderInfo_isExplicit(x_65); -if (x_66 == 0) -{ -lean_object* x_67; -x_67 = lean_box(0); -x_33 = x_67; -goto block_62; + lean_dec_ref(x_26); + x_29 = lean_box(0); } -else +x_67 = lean_unbox(x_27); +lean_dec(x_27); +if (x_67 == 0) { lean_object* x_68; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); -x_68 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); +x_68 = lean_ctor_get(x_2, 1); +lean_inc(x_68); if (lean_obj_tag(x_68) == 0) { -uint8_t x_69; -x_69 = !lean_is_exclusive(x_68); -if (x_69 == 0) -{ -lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; lean_object* x_74; -x_70 = lean_ctor_get(x_68, 0); -x_71 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_71, 0, x_70); -x_72 = 1; -x_73 = lean_box(x_72); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_71); -lean_ctor_set(x_68, 0, x_74); -return x_68; +lean_object* x_69; +lean_free_object(x_19); +x_69 = lean_box(0); +x_30 = x_69; +goto block_66; } else { -lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_75 = lean_ctor_get(x_68, 0); -x_76 = lean_ctor_get(x_68, 1); -lean_inc(x_76); -lean_inc(x_75); +lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_70 = lean_ctor_get(x_68, 0); +lean_inc(x_70); lean_dec(x_68); -x_77 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_77, 0, x_75); -x_78 = 1; -x_79 = lean_box(x_78); -x_80 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_80, 0, x_79); -lean_ctor_set(x_80, 1, x_77); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_76); -return x_81; -} +x_71 = lean_unsigned_to_nat(0u); +x_72 = lean_nat_dec_eq(x_1, x_71); +if (x_72 == 0) +{ +lean_object* x_73; +lean_dec(x_70); +lean_free_object(x_19); +x_73 = lean_box(0); +x_30 = x_73; +goto block_66; } else { -uint8_t x_82; -x_82 = !lean_is_exclusive(x_68); -if (x_82 == 0) +uint8_t x_74; +x_74 = lean_expr_eqv(x_70, x_11); +lean_dec(x_70); +if (x_74 == 0) { -return x_68; +lean_object* x_75; +lean_free_object(x_19); +x_75 = lean_box(0); +x_30 = x_75; +goto block_66; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_68, 0); -x_84 = lean_ctor_get(x_68, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_68); -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_76; +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_76 = lean_box(0); +lean_ctor_set(x_19, 1, x_28); +lean_ctor_set(x_19, 0, x_76); +return x_19; } } } } else { -uint8_t x_86; -x_86 = !lean_is_exclusive(x_64); -if (x_86 == 0) +lean_object* x_77; lean_object* x_78; +lean_dec(x_29); +lean_free_object(x_19); +lean_dec(x_11); +x_77 = lean_ctor_get(x_2, 0); +lean_inc(x_77); +lean_dec(x_2); +x_78 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_77, x_3, x_4, x_5, x_6, x_7, x_8, x_28); +return x_78; +} +block_66: { -lean_object* x_87; lean_object* x_88; uint8_t x_89; -x_87 = lean_ctor_get(x_64, 0); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_nat_dec_eq(x_1, x_88); -lean_dec(x_1); -if (x_89 == 0) +uint8_t x_31; uint8_t x_32; +lean_dec(x_30); +x_31 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_32 = l_Lean_BinderInfo_isExplicit(x_31); +if (x_32 == 0) { -uint8_t x_90; uint8_t x_91; -lean_dec(x_87); -lean_free_object(x_20); -x_90 = lean_ctor_get_uint8(x_27, sizeof(void*)*2); -x_91 = l_Lean_BinderInfo_isExplicit(x_90); -if (x_91 == 0) +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_33 = lean_ctor_get(x_7, 2); +lean_inc(x_33); +x_34 = lean_ctor_get(x_2, 0); +lean_inc(x_34); +lean_dec(x_2); +x_35 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4; +x_36 = lean_name_eq(x_34, x_35); +if (x_36 == 0) { -lean_object* x_92; -lean_free_object(x_64); -x_92 = lean_box(0); -x_33 = x_92; -goto block_62; +lean_object* x_37; lean_object* x_38; +lean_dec(x_34); +lean_dec(x_33); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_37 = lean_box(0); +if (lean_is_scalar(x_29)) { + x_38 = lean_alloc_ctor(0, 2, 0); +} else { + x_38 = x_29; +} +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_28); +return x_38; } else { -lean_object* x_93; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); -x_93 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_93) == 0) +lean_object* x_39; +lean_dec(x_29); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_39 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_11, x_33, x_5, x_6, x_7, x_8, x_28); +lean_dec(x_33); +if (lean_obj_tag(x_39) == 0) { -uint8_t x_94; -x_94 = !lean_is_exclusive(x_93); -if (x_94 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_39, 0); +lean_inc(x_40); +x_41 = lean_unbox(x_40); +lean_dec(x_40); +if (x_41 == 0) { -lean_object* x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; -x_95 = lean_ctor_get(x_93, 0); -lean_ctor_set(x_64, 0, x_95); -x_96 = 1; -x_97 = lean_box(x_96); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_64); -lean_ctor_set(x_93, 0, x_98); -return x_93; +uint8_t x_42; +lean_dec(x_34); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_42 = !lean_is_exclusive(x_39); +if (x_42 == 0) +{ +lean_object* x_43; lean_object* x_44; +x_43 = lean_ctor_get(x_39, 0); +lean_dec(x_43); +x_44 = lean_box(0); +lean_ctor_set(x_39, 0, x_44); +return x_39; } else { -lean_object* x_99; lean_object* x_100; uint8_t x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_99 = lean_ctor_get(x_93, 0); -x_100 = lean_ctor_get(x_93, 1); -lean_inc(x_100); -lean_inc(x_99); -lean_dec(x_93); -lean_ctor_set(x_64, 0, x_99); -x_101 = 1; -x_102 = lean_box(x_101); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_64); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_100); -return x_104; +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_39, 1); +lean_inc(x_45); +lean_dec(x_39); +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 { -uint8_t x_105; -lean_free_object(x_64); -x_105 = !lean_is_exclusive(x_93); -if (x_105 == 0) -{ -return x_93; +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_39, 1); +lean_inc(x_48); +lean_dec(x_39); +x_49 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_34, x_3, x_4, x_5, x_6, x_7, x_8, x_48); +return x_49; +} } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_93, 0); -x_107 = lean_ctor_get(x_93, 1); -lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_93); -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; +uint8_t x_50; +lean_dec(x_34); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_50 = !lean_is_exclusive(x_39); +if (x_50 == 0) +{ +return x_39; } +else +{ +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_39, 0); +x_52 = lean_ctor_get(x_39, 1); +lean_inc(x_52); +lean_inc(x_51); +lean_dec(x_39); +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 -{ -uint8_t x_109; -x_109 = lean_expr_eqv(x_87, x_25); -lean_dec(x_87); -if (x_109 == 0) -{ -uint8_t x_110; uint8_t x_111; -lean_free_object(x_20); -x_110 = lean_ctor_get_uint8(x_27, sizeof(void*)*2); -x_111 = l_Lean_BinderInfo_isExplicit(x_110); -if (x_111 == 0) -{ -lean_object* x_112; -lean_free_object(x_64); -x_112 = lean_box(0); -x_33 = x_112; -goto block_62; } else { -lean_object* x_113; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); -x_113 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_113) == 0) +lean_object* x_54; +lean_dec(x_29); +lean_dec(x_11); +lean_dec(x_2); +x_54 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_28); +if (lean_obj_tag(x_54) == 0) { -uint8_t x_114; -x_114 = !lean_is_exclusive(x_113); -if (x_114 == 0) +uint8_t x_55; +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_115; uint8_t x_116; lean_object* x_117; lean_object* x_118; -x_115 = lean_ctor_get(x_113, 0); -lean_ctor_set(x_64, 0, x_115); -x_116 = 1; -x_117 = lean_box(x_116); -x_118 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_64); -lean_ctor_set(x_113, 0, x_118); -return x_113; +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_54, 0, x_57); +return x_54; } else { -lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -x_119 = lean_ctor_get(x_113, 0); -x_120 = lean_ctor_get(x_113, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_113); -lean_ctor_set(x_64, 0, x_119); -x_121 = 1; -x_122 = lean_box(x_121); -x_123 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_64); -x_124 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_124, 0, x_123); -lean_ctor_set(x_124, 1, x_120); -return x_124; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_58 = lean_ctor_get(x_54, 0); +x_59 = lean_ctor_get(x_54, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_54); +x_60 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_60, 0, x_58); +x_61 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_61, 0, x_60); +lean_ctor_set(x_61, 1, x_59); +return x_61; } } else { -uint8_t x_125; -lean_free_object(x_64); -x_125 = !lean_is_exclusive(x_113); -if (x_125 == 0) +uint8_t x_62; +x_62 = !lean_is_exclusive(x_54); +if (x_62 == 0) { -return x_113; +return x_54; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_113, 0); -x_127 = lean_ctor_get(x_113, 1); -lean_inc(x_127); -lean_inc(x_126); -lean_dec(x_113); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_54, 0); +x_64 = lean_ctor_get(x_54, 1); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_54); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_64); +return x_65; +} } } } } else { -lean_object* x_129; -lean_free_object(x_64); -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); +uint8_t x_79; +lean_free_object(x_19); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_129 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -lean_ctor_set(x_20, 1, x_31); -lean_ctor_set(x_20, 0, x_129); -return x_20; -} -} -} -else -{ -lean_object* x_130; lean_object* x_131; uint8_t x_132; -x_130 = lean_ctor_get(x_64, 0); -lean_inc(x_130); -lean_dec(x_64); -x_131 = lean_unsigned_to_nat(0u); -x_132 = lean_nat_dec_eq(x_1, x_131); -lean_dec(x_1); -if (x_132 == 0) -{ -uint8_t x_133; uint8_t x_134; -lean_dec(x_130); -lean_free_object(x_20); -x_133 = lean_ctor_get_uint8(x_27, sizeof(void*)*2); -x_134 = l_Lean_BinderInfo_isExplicit(x_133); -if (x_134 == 0) +lean_dec(x_2); +x_79 = !lean_is_exclusive(x_26); +if (x_79 == 0) { -lean_object* x_135; -x_135 = lean_box(0); -x_33 = x_135; -goto block_62; +return x_26; } else { -lean_object* x_136; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); -x_136 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_136) == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_139 = x_136; -} else { - lean_dec_ref(x_136); - x_139 = lean_box(0); -} -x_140 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_140, 0, x_137); -x_141 = 1; -x_142 = lean_box(x_141); -x_143 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_143, 0, x_142); -lean_ctor_set(x_143, 1, x_140); -if (lean_is_scalar(x_139)) { - x_144 = lean_alloc_ctor(0, 2, 0); -} else { - x_144 = x_139; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_26, 0); +x_81 = lean_ctor_get(x_26, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_26); +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_ctor_set(x_144, 0, x_143); -lean_ctor_set(x_144, 1, x_138); -return x_144; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; -x_145 = lean_ctor_get(x_136, 0); -lean_inc(x_145); -x_146 = lean_ctor_get(x_136, 1); -lean_inc(x_146); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_147 = x_136; -} else { - lean_dec_ref(x_136); - x_147 = lean_box(0); -} -if (lean_is_scalar(x_147)) { - x_148 = lean_alloc_ctor(1, 2, 0); -} else { - x_148 = x_147; -} -lean_ctor_set(x_148, 0, x_145); -lean_ctor_set(x_148, 1, x_146); -return x_148; -} +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_19, 1); +lean_inc(x_83); +lean_dec(x_19); +x_84 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_85 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_84, x_3, x_4, x_5, x_6, x_7, x_8, x_83); +if (lean_obj_tag(x_85) == 0) +{ +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_122; +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; +} else { + lean_dec_ref(x_85); + x_88 = lean_box(0); } +x_122 = lean_unbox(x_86); +lean_dec(x_86); +if (x_122 == 0) +{ +lean_object* x_123; +x_123 = lean_ctor_get(x_2, 1); +lean_inc(x_123); +if (lean_obj_tag(x_123) == 0) +{ +lean_object* x_124; +x_124 = lean_box(0); +x_89 = x_124; +goto block_121; } else { -uint8_t x_149; -x_149 = lean_expr_eqv(x_130, x_25); -lean_dec(x_130); -if (x_149 == 0) -{ -uint8_t x_150; uint8_t x_151; -lean_free_object(x_20); -x_150 = lean_ctor_get_uint8(x_27, sizeof(void*)*2); -x_151 = l_Lean_BinderInfo_isExplicit(x_150); -if (x_151 == 0) +lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_125 = lean_ctor_get(x_123, 0); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_unsigned_to_nat(0u); +x_127 = lean_nat_dec_eq(x_1, x_126); +if (x_127 == 0) { -lean_object* x_152; -x_152 = lean_box(0); -x_33 = x_152; -goto block_62; +lean_object* x_128; +lean_dec(x_125); +x_128 = lean_box(0); +x_89 = x_128; +goto block_121; } else { -lean_object* x_153; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); -x_153 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_153) == 0) +uint8_t x_129; +x_129 = lean_expr_eqv(x_125, x_11); +lean_dec(x_125); +if (x_129 == 0) { -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; uint8_t x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_154 = lean_ctor_get(x_153, 0); -lean_inc(x_154); -x_155 = lean_ctor_get(x_153, 1); -lean_inc(x_155); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_156 = x_153; -} else { - lean_dec_ref(x_153); - x_156 = lean_box(0); -} -x_157 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_157, 0, x_154); -x_158 = 1; -x_159 = lean_box(x_158); -x_160 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_157); -if (lean_is_scalar(x_156)) { - x_161 = lean_alloc_ctor(0, 2, 0); -} else { - x_161 = x_156; -} -lean_ctor_set(x_161, 0, x_160); -lean_ctor_set(x_161, 1, x_155); -return x_161; +lean_object* x_130; +x_130 = lean_box(0); +x_89 = x_130; +goto block_121; } else { -lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_162 = lean_ctor_get(x_153, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_153, 1); -lean_inc(x_163); -if (lean_is_exclusive(x_153)) { - lean_ctor_release(x_153, 0); - lean_ctor_release(x_153, 1); - x_164 = x_153; -} else { - lean_dec_ref(x_153); - x_164 = lean_box(0); -} -if (lean_is_scalar(x_164)) { - x_165 = lean_alloc_ctor(1, 2, 0); -} else { - x_165 = x_164; +lean_object* x_131; lean_object* x_132; +lean_dec(x_88); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_131 = lean_box(0); +x_132 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_87); +return x_132; } -lean_ctor_set(x_165, 0, x_162); -lean_ctor_set(x_165, 1, x_163); -return x_165; } } } else { -lean_object* x_166; -lean_dec(x_32); -lean_dec(x_27); -lean_dec(x_25); +lean_object* x_133; lean_object* x_134; +lean_dec(x_88); +lean_dec(x_11); +x_133 = lean_ctor_get(x_2, 0); +lean_inc(x_133); +lean_dec(x_2); +x_134 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_133, x_3, x_4, x_5, x_6, x_7, x_8, x_87); +return x_134; +} +block_121: +{ +uint8_t x_90; uint8_t x_91; +lean_dec(x_89); +x_90 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +x_91 = l_Lean_BinderInfo_isExplicit(x_90); +if (x_91 == 0) +{ +lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; +x_92 = lean_ctor_get(x_7, 2); +lean_inc(x_92); +x_93 = lean_ctor_get(x_2, 0); +lean_inc(x_93); +lean_dec(x_2); +x_94 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4; +x_95 = lean_name_eq(x_93, x_94); +if (x_95 == 0) +{ +lean_object* x_96; lean_object* x_97; +lean_dec(x_93); +lean_dec(x_92); +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_166 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -lean_ctor_set(x_20, 1, x_31); -lean_ctor_set(x_20, 0, x_166); -return x_20; -} -} -} +x_96 = lean_box(0); +if (lean_is_scalar(x_88)) { + x_97 = lean_alloc_ctor(0, 2, 0); +} else { + x_97 = x_88; } +lean_ctor_set(x_97, 0, x_96); +lean_ctor_set(x_97, 1, x_87); +return x_97; } else { -lean_object* x_167; -lean_dec(x_32); -lean_free_object(x_20); -lean_dec(x_25); -lean_dec(x_1); +lean_object* x_98; +lean_dec(x_88); 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_167 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_31); -if (lean_obj_tag(x_167) == 0) +x_98 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_11, x_92, x_5, x_6, x_7, x_8, x_87); +lean_dec(x_92); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_168 = lean_ctor_get(x_167, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_167, 1); -lean_inc(x_169); -lean_dec(x_167); -x_170 = lean_ctor_get(x_27, 0); -lean_inc(x_170); -lean_dec(x_27); -x_171 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_170, x_168, x_3, x_4, x_5, x_6, x_7, x_8, x_169); +lean_object* x_99; uint8_t x_100; +x_99 = lean_ctor_get(x_98, 0); +lean_inc(x_99); +x_100 = lean_unbox(x_99); +lean_dec(x_99); +if (x_100 == 0) +{ +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +lean_dec(x_93); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_171; +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; +} else { + lean_dec_ref(x_98); + x_102 = lean_box(0); +} +x_103 = lean_box(0); +if (lean_is_scalar(x_102)) { + x_104 = lean_alloc_ctor(0, 2, 0); +} else { + x_104 = x_102; +} +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +return x_104; } else { -uint8_t x_172; -lean_dec(x_27); +lean_object* x_105; lean_object* x_106; +x_105 = lean_ctor_get(x_98, 1); +lean_inc(x_105); +lean_dec(x_98); +x_106 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_93, x_3, x_4, x_5, x_6, x_7, x_8, x_105); +return x_106; +} +} +else +{ +lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; +lean_dec(x_93); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_172 = !lean_is_exclusive(x_167); -if (x_172 == 0) -{ -return x_167; +x_107 = lean_ctor_get(x_98, 0); +lean_inc(x_107); +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); +} +if (lean_is_scalar(x_109)) { + x_110 = lean_alloc_ctor(1, 2, 0); +} else { + x_110 = x_109; +} +lean_ctor_set(x_110, 0, x_107); +lean_ctor_set(x_110, 1, x_108); +return x_110; +} +} } else { -lean_object* x_173; lean_object* x_174; lean_object* x_175; -x_173 = lean_ctor_get(x_167, 0); -x_174 = lean_ctor_get(x_167, 1); -lean_inc(x_174); -lean_inc(x_173); -lean_dec(x_167); -x_175 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_175, 0, x_173); -lean_ctor_set(x_175, 1, x_174); -return x_175; +lean_object* x_111; +lean_dec(x_88); +lean_dec(x_11); +lean_dec(x_2); +x_111 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_87); +if (lean_obj_tag(x_111) == 0) +{ +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_112 = lean_ctor_get(x_111, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_111, 1); +lean_inc(x_113); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_114 = x_111; +} else { + lean_dec_ref(x_111); + x_114 = lean_box(0); } +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_112); +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; } -block_62: +else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; -lean_dec(x_33); -x_34 = lean_ctor_get(x_7, 2); -lean_inc(x_34); -x_35 = lean_ctor_get(x_27, 0); -lean_inc(x_35); -lean_dec(x_27); -x_36 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6; -x_37 = lean_name_eq(x_35, x_36); -if (x_37 == 0) +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_117 = lean_ctor_get(x_111, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_111, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_111)) { + lean_ctor_release(x_111, 0); + lean_ctor_release(x_111, 1); + x_119 = x_111; +} else { + lean_dec_ref(x_111); + x_119 = lean_box(0); +} +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); +} else { + x_120 = x_119; +} +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; +} +} +} +} +else { -lean_object* x_38; lean_object* x_39; -lean_dec(x_35); -lean_dec(x_34); -lean_dec(x_25); +lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_38 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -if (lean_is_scalar(x_32)) { - x_39 = lean_alloc_ctor(0, 2, 0); +lean_dec(x_2); +x_135 = lean_ctor_get(x_85, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_85, 1); +lean_inc(x_136); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_137 = x_85; } else { - x_39 = x_32; + lean_dec_ref(x_85); + x_137 = lean_box(0); +} +if (lean_is_scalar(x_137)) { + x_138 = lean_alloc_ctor(1, 2, 0); +} else { + x_138 = x_137; +} +lean_ctor_set(x_138, 0, x_135); +lean_ctor_set(x_138, 1, x_136); +return x_138; +} } -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_31); -return x_39; } else { -lean_object* x_40; -lean_dec(x_32); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_40 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_25, x_34, x_5, x_6, x_7, x_8, x_31); -lean_dec(x_34); -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) -{ -uint8_t x_43; -lean_dec(x_35); +uint8_t x_139; +lean_dec(x_11); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_43 = !lean_is_exclusive(x_40); -if (x_43 == 0) +lean_dec(x_2); +x_139 = !lean_is_exclusive(x_19); +if (x_139 == 0) { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_40, 0); -lean_dec(x_44); -x_45 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -lean_ctor_set(x_40, 0, x_45); -return x_40; +lean_object* x_140; lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_140 = lean_ctor_get(x_19, 0); +lean_dec(x_140); +x_141 = lean_ctor_get(x_7, 5); +lean_inc(x_141); +lean_dec(x_7); +x_142 = 0; +x_143 = l_Lean_SourceInfo_fromRef(x_141, x_142); +x_144 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__6; +lean_inc(x_143); +x_145 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_145, 0, x_143); +lean_ctor_set(x_145, 1, x_144); +x_146 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__5; +x_147 = l_Lean_Syntax_node1(x_143, x_146, x_145); +x_148 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_148, 0, x_147); +lean_ctor_set(x_19, 0, x_148); +return x_19; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_40, 1); -lean_inc(x_46); -lean_dec(x_40); -x_47 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -x_48 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_46); -return x_48; +lean_object* x_149; lean_object* x_150; uint8_t x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; +x_149 = lean_ctor_get(x_19, 1); +lean_inc(x_149); +lean_dec(x_19); +x_150 = lean_ctor_get(x_7, 5); +lean_inc(x_150); +lean_dec(x_7); +x_151 = 0; +x_152 = l_Lean_SourceInfo_fromRef(x_150, x_151); +x_153 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__6; +lean_inc(x_152); +x_154 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +x_155 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__5; +x_156 = l_Lean_Syntax_node1(x_152, x_155, x_154); +x_157 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_157, 0, x_156); +x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_149); +return x_158; } } -else -{ -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_40, 1); -lean_inc(x_49); -lean_dec(x_40); -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_50 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_49); -if (lean_obj_tag(x_50) == 0) -{ -lean_object* x_51; lean_object* x_52; lean_object* x_53; -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_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_35, x_51, x_3, x_4, x_5, x_6, x_7, x_8, x_52); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_53; } else { -uint8_t x_54; -lean_dec(x_35); +uint8_t x_159; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_54 = !lean_is_exclusive(x_50); -if (x_54 == 0) +lean_dec(x_2); +x_159 = !lean_is_exclusive(x_19); +if (x_159 == 0) { -return x_50; +return x_19; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_50, 0); -x_56 = lean_ctor_get(x_50, 1); -lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_50); -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_160; lean_object* x_161; lean_object* x_162; +x_160 = lean_ctor_get(x_19, 0); +x_161 = lean_ctor_get(x_19, 1); +lean_inc(x_161); +lean_inc(x_160); +lean_dec(x_19); +x_162 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_162, 0, x_160); +lean_ctor_set(x_162, 1, x_161); +return x_162; } } } else { -uint8_t x_58; -lean_dec(x_35); +uint8_t x_163; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_58 = !lean_is_exclusive(x_40); -if (x_58 == 0) +lean_dec(x_2); +x_163 = !lean_is_exclusive(x_14); +if (x_163 == 0) { -return x_40; +lean_object* x_164; lean_object* x_165; +x_164 = lean_ctor_get(x_14, 0); +lean_dec(x_164); +x_165 = lean_box(0); +lean_ctor_set(x_14, 0, x_165); +return x_14; } else { -lean_object* x_59; lean_object* x_60; lean_object* x_61; -x_59 = lean_ctor_get(x_40, 0); -x_60 = lean_ctor_get(x_40, 1); -lean_inc(x_60); -lean_inc(x_59); -lean_dec(x_40); -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_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_14, 1); +lean_inc(x_166); +lean_dec(x_14); +x_167 = lean_box(0); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_166); +return x_168; } } } else { -uint8_t x_176; -lean_dec(x_27); -lean_free_object(x_20); -lean_dec(x_25); +uint8_t x_169; +lean_dec(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -x_176 = !lean_is_exclusive(x_29); -if (x_176 == 0) +lean_dec(x_2); +x_169 = !lean_is_exclusive(x_14); +if (x_169 == 0) { -return x_29; +return x_14; } else { -lean_object* x_177; lean_object* x_178; lean_object* x_179; -x_177 = lean_ctor_get(x_29, 0); -x_178 = lean_ctor_get(x_29, 1); -lean_inc(x_178); -lean_inc(x_177); -lean_dec(x_29); -x_179 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_179, 0, x_177); -lean_ctor_set(x_179, 1, x_178); -return x_179; +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_14, 0); +x_171 = lean_ctor_get(x_14, 1); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_14); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +return x_172; } } } -else +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; -x_180 = lean_ctor_get(x_20, 0); -x_181 = lean_ctor_get(x_20, 1); -lean_inc(x_181); -lean_inc(x_180); -lean_dec(x_20); -x_182 = lean_ctor_get(x_2, 0); -lean_inc(x_182); -lean_dec(x_2); -x_183 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4; -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_184 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_183, x_3, x_4, x_5, x_6, x_7, x_8, x_181); -if (lean_obj_tag(x_184) == 0) +lean_object* x_10; +x_10 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +static lean_object* _init_l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1() { +_start: { -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; uint8_t x_216; -x_185 = lean_ctor_get(x_184, 0); -lean_inc(x_185); -x_186 = lean_ctor_get(x_184, 1); -lean_inc(x_186); -if (lean_is_exclusive(x_184)) { - lean_ctor_release(x_184, 0); - lean_ctor_release(x_184, 1); - x_187 = x_184; -} else { - lean_dec_ref(x_184); - x_187 = lean_box(0); +lean_object* x_1; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabFailureId; +return x_1; } -x_216 = lean_unbox(x_185); -lean_dec(x_185); -if (x_216 == 0) +} +LEAN_EXPORT lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_217; -x_217 = lean_ctor_get(x_182, 1); -lean_inc(x_217); -if (lean_obj_tag(x_217) == 0) +if (lean_obj_tag(x_3) == 0) { -uint8_t x_218; uint8_t x_219; +lean_object* x_11; +lean_dec(x_2); lean_dec(x_1); -x_218 = lean_ctor_get_uint8(x_182, sizeof(void*)*2); -x_219 = l_Lean_BinderInfo_isExplicit(x_218); -if (x_219 == 0) -{ -lean_object* x_220; -x_220 = lean_box(0); -x_188 = x_220; -goto block_215; +x_11 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_10); +return x_11; } else { -lean_object* x_221; -lean_dec(x_187); -lean_dec(x_182); -lean_dec(x_180); -x_221 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_186); -if (lean_obj_tag(x_221) == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = lean_ctor_get(x_3, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_3, 1); +lean_inc(x_13); +lean_dec(x_3); +x_14 = lean_box(0); +lean_inc(x_2); +lean_inc(x_1); +x_15 = lean_apply_3(x_12, x_1, x_2, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; uint8_t x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; -x_222 = lean_ctor_get(x_221, 0); -lean_inc(x_222); -x_223 = lean_ctor_get(x_221, 1); -lean_inc(x_223); -if (lean_is_exclusive(x_221)) { - lean_ctor_release(x_221, 0); - lean_ctor_release(x_221, 1); - x_224 = x_221; -} else { - lean_dec_ref(x_221); - x_224 = lean_box(0); -} -x_225 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_225, 0, x_222); -x_226 = 1; -x_227 = lean_box(x_226); -x_228 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_228, 0, x_227); -lean_ctor_set(x_228, 1, x_225); -if (lean_is_scalar(x_224)) { - x_229 = lean_alloc_ctor(0, 2, 0); -} else { - x_229 = x_224; -} -lean_ctor_set(x_229, 0, x_228); -lean_ctor_set(x_229, 1, x_223); -return x_229; +uint8_t x_16; +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 1); +lean_dec(x_17); +lean_ctor_set(x_15, 1, x_10); +return x_15; } else { -lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; -x_230 = lean_ctor_get(x_221, 0); -lean_inc(x_230); -x_231 = lean_ctor_get(x_221, 1); -lean_inc(x_231); -if (lean_is_exclusive(x_221)) { - lean_ctor_release(x_221, 0); - lean_ctor_release(x_221, 1); - x_232 = x_221; -} else { - lean_dec_ref(x_221); - x_232 = lean_box(0); -} -if (lean_is_scalar(x_232)) { - x_233 = lean_alloc_ctor(1, 2, 0); -} else { - x_233 = x_232; -} -lean_ctor_set(x_233, 0, x_230); -lean_ctor_set(x_233, 1, x_231); -return x_233; -} +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_inc(x_18); +lean_dec(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_10); +return x_19; } } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; uint8_t x_237; -x_234 = lean_ctor_get(x_217, 0); -lean_inc(x_234); -if (lean_is_exclusive(x_217)) { - lean_ctor_release(x_217, 0); - x_235 = x_217; -} else { - lean_dec_ref(x_217); - x_235 = lean_box(0); -} -x_236 = lean_unsigned_to_nat(0u); -x_237 = lean_nat_dec_eq(x_1, x_236); -lean_dec(x_1); -if (x_237 == 0) +lean_object* x_20; uint8_t x_21; +lean_dec(x_15); +x_20 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_10); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -uint8_t x_238; uint8_t x_239; -lean_dec(x_234); -x_238 = lean_ctor_get_uint8(x_182, sizeof(void*)*2); -x_239 = l_Lean_BinderInfo_isExplicit(x_238); -if (x_239 == 0) +lean_object* x_22; lean_object* x_23; uint8_t x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = l_Lean_Exception_isRuntime(x_22); +if (x_24 == 0) { -lean_object* x_240; -lean_dec(x_235); -x_240 = lean_box(0); -x_188 = x_240; -goto block_215; -} -else +lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_25 = lean_ctor_get(x_22, 0); +lean_inc(x_25); +x_26 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_27 = lean_nat_dec_eq(x_26, x_25); +lean_dec(x_25); +if (x_27 == 0) { -lean_object* x_241; -lean_dec(x_187); -lean_dec(x_182); -lean_dec(x_180); -x_241 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_186); -if (lean_obj_tag(x_241) == 0) -{ -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; uint8_t x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; -x_242 = lean_ctor_get(x_241, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_241, 1); -lean_inc(x_243); -if (lean_is_exclusive(x_241)) { - lean_ctor_release(x_241, 0); - lean_ctor_release(x_241, 1); - x_244 = x_241; -} else { - lean_dec_ref(x_241); - x_244 = lean_box(0); -} -if (lean_is_scalar(x_235)) { - x_245 = lean_alloc_ctor(1, 1, 0); -} else { - x_245 = x_235; -} -lean_ctor_set(x_245, 0, x_242); -x_246 = 1; -x_247 = lean_box(x_246); -x_248 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_248, 0, x_247); -lean_ctor_set(x_248, 1, x_245); -if (lean_is_scalar(x_244)) { - x_249 = lean_alloc_ctor(0, 2, 0); -} else { - x_249 = x_244; -} -lean_ctor_set(x_249, 0, x_248); -lean_ctor_set(x_249, 1, x_243); -return x_249; +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +return x_20; } else { -lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; -lean_dec(x_235); -x_250 = lean_ctor_get(x_241, 0); -lean_inc(x_250); -x_251 = lean_ctor_get(x_241, 1); -lean_inc(x_251); -if (lean_is_exclusive(x_241)) { - lean_ctor_release(x_241, 0); - lean_ctor_release(x_241, 1); - x_252 = x_241; -} else { - lean_dec_ref(x_241); - x_252 = lean_box(0); -} -if (lean_is_scalar(x_252)) { - x_253 = lean_alloc_ctor(1, 2, 0); -} else { - x_253 = x_252; -} -lean_ctor_set(x_253, 0, x_250); -lean_ctor_set(x_253, 1, x_251); -return x_253; -} +lean_free_object(x_20); +lean_dec(x_22); +x_3 = x_13; +x_10 = x_23; +goto _start; } } else { -uint8_t x_254; -x_254 = lean_expr_eqv(x_234, x_180); -lean_dec(x_234); -if (x_254 == 0) -{ -uint8_t x_255; uint8_t x_256; -x_255 = lean_ctor_get_uint8(x_182, sizeof(void*)*2); -x_256 = l_Lean_BinderInfo_isExplicit(x_255); -if (x_256 == 0) +uint8_t x_29; +x_29 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +if (x_29 == 0) { -lean_object* x_257; -lean_dec(x_235); -x_257 = lean_box(0); -x_188 = x_257; -goto block_215; +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +return x_20; } else { -lean_object* x_258; -lean_dec(x_187); -lean_dec(x_182); -lean_dec(x_180); -x_258 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_186); -if (lean_obj_tag(x_258) == 0) +lean_object* x_30; lean_object* x_31; uint8_t x_32; +x_30 = lean_ctor_get(x_22, 0); +lean_inc(x_30); +x_31 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_32 = lean_nat_dec_eq(x_31, x_30); +lean_dec(x_30); +if (x_32 == 0) { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; uint8_t x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_258, 1); -lean_inc(x_260); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_261 = x_258; -} else { - lean_dec_ref(x_258); - x_261 = lean_box(0); -} -if (lean_is_scalar(x_235)) { - x_262 = lean_alloc_ctor(1, 1, 0); -} else { - x_262 = x_235; -} -lean_ctor_set(x_262, 0, x_259); -x_263 = 1; -x_264 = lean_box(x_263); -x_265 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_265, 0, x_264); -lean_ctor_set(x_265, 1, x_262); -if (lean_is_scalar(x_261)) { - x_266 = lean_alloc_ctor(0, 2, 0); -} else { - x_266 = x_261; -} -lean_ctor_set(x_266, 0, x_265); -lean_ctor_set(x_266, 1, x_260); -return x_266; +lean_dec(x_13); +lean_dec(x_2); +lean_dec(x_1); +return x_20; } else { -lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -lean_dec(x_235); -x_267 = lean_ctor_get(x_258, 0); -lean_inc(x_267); -x_268 = lean_ctor_get(x_258, 1); -lean_inc(x_268); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_269 = x_258; -} else { - lean_dec_ref(x_258); - x_269 = lean_box(0); -} -if (lean_is_scalar(x_269)) { - x_270 = lean_alloc_ctor(1, 2, 0); -} else { - x_270 = x_269; -} -lean_ctor_set(x_270, 0, x_267); -lean_ctor_set(x_270, 1, x_268); -return x_270; -} -} -} -else -{ -lean_object* x_271; lean_object* x_272; -lean_dec(x_235); -lean_dec(x_187); -lean_dec(x_182); -lean_dec(x_180); -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_271 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -x_272 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_186); -return x_272; -} -} -} -} -else -{ -lean_object* x_273; -lean_dec(x_187); -lean_dec(x_180); -lean_dec(x_1); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_273 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_186); -if (lean_obj_tag(x_273) == 0) -{ -lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -x_274 = lean_ctor_get(x_273, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_273, 1); -lean_inc(x_275); -lean_dec(x_273); -x_276 = lean_ctor_get(x_182, 0); -lean_inc(x_276); -lean_dec(x_182); -x_277 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_276, x_274, x_3, x_4, x_5, x_6, x_7, x_8, x_275); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_277; -} -else -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; -lean_dec(x_182); -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_278 = lean_ctor_get(x_273, 0); -lean_inc(x_278); -x_279 = lean_ctor_get(x_273, 1); -lean_inc(x_279); -if (lean_is_exclusive(x_273)) { - lean_ctor_release(x_273, 0); - lean_ctor_release(x_273, 1); - x_280 = x_273; -} else { - lean_dec_ref(x_273); - 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; -} -} -block_215: -{ -lean_object* x_189; lean_object* x_190; lean_object* x_191; uint8_t x_192; -lean_dec(x_188); -x_189 = lean_ctor_get(x_7, 2); -lean_inc(x_189); -x_190 = lean_ctor_get(x_182, 0); -lean_inc(x_190); -lean_dec(x_182); -x_191 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6; -x_192 = lean_name_eq(x_190, x_191); -if (x_192 == 0) -{ -lean_object* x_193; lean_object* x_194; -lean_dec(x_190); -lean_dec(x_189); -lean_dec(x_180); -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_193 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -if (lean_is_scalar(x_187)) { - x_194 = lean_alloc_ctor(0, 2, 0); -} else { - x_194 = x_187; -} -lean_ctor_set(x_194, 0, x_193); -lean_ctor_set(x_194, 1, x_186); -return x_194; -} -else -{ -lean_object* x_195; -lean_dec(x_187); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_195 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_180, x_189, x_5, x_6, x_7, x_8, x_186); -lean_dec(x_189); -if (lean_obj_tag(x_195) == 0) -{ -lean_object* x_196; uint8_t x_197; -x_196 = lean_ctor_get(x_195, 0); -lean_inc(x_196); -x_197 = lean_unbox(x_196); -lean_dec(x_196); -if (x_197 == 0) -{ -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -lean_dec(x_190); -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_198 = lean_ctor_get(x_195, 1); -lean_inc(x_198); -if (lean_is_exclusive(x_195)) { - lean_ctor_release(x_195, 0); - lean_ctor_release(x_195, 1); - x_199 = x_195; -} else { - lean_dec_ref(x_195); - x_199 = lean_box(0); -} -x_200 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -if (lean_is_scalar(x_199)) { - x_201 = lean_alloc_ctor(0, 2, 0); -} else { - x_201 = x_199; -} -lean_ctor_set(x_201, 0, x_200); -lean_ctor_set(x_201, 1, x_198); -return x_201; -} -else -{ -lean_object* x_202; lean_object* x_203; -x_202 = lean_ctor_get(x_195, 1); -lean_inc(x_202); -lean_dec(x_195); -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_203 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_7, x_8, x_202); -if (lean_obj_tag(x_203) == 0) -{ -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_ctor_get(x_203, 0); -lean_inc(x_204); -x_205 = lean_ctor_get(x_203, 1); -lean_inc(x_205); -lean_dec(x_203); -x_206 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg(x_190, x_204, x_3, x_4, x_5, x_6, x_7, x_8, x_205); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_206; -} -else -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; -lean_dec(x_190); -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_207 = lean_ctor_get(x_203, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_203, 1); -lean_inc(x_208); -if (lean_is_exclusive(x_203)) { - lean_ctor_release(x_203, 0); - lean_ctor_release(x_203, 1); - x_209 = x_203; -} else { - lean_dec_ref(x_203); - x_209 = lean_box(0); -} -if (lean_is_scalar(x_209)) { - x_210 = lean_alloc_ctor(1, 2, 0); -} else { - x_210 = x_209; -} -lean_ctor_set(x_210, 0, x_207); -lean_ctor_set(x_210, 1, x_208); -return x_210; -} -} -} -else -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_190); -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_211 = lean_ctor_get(x_195, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_195, 1); -lean_inc(x_212); -if (lean_is_exclusive(x_195)) { - lean_ctor_release(x_195, 0); - lean_ctor_release(x_195, 1); - x_213 = x_195; -} else { - lean_dec_ref(x_195); - x_213 = lean_box(0); -} -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(1, 2, 0); -} else { - x_214 = x_213; -} -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set(x_214, 1, x_212); -return x_214; -} -} -} -} -else -{ -lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; -lean_dec(x_182); -lean_dec(x_180); -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_282 = lean_ctor_get(x_184, 0); -lean_inc(x_282); -x_283 = lean_ctor_get(x_184, 1); -lean_inc(x_283); -if (lean_is_exclusive(x_184)) { - lean_ctor_release(x_184, 0); - lean_ctor_release(x_184, 1); - x_284 = x_184; -} else { - lean_dec_ref(x_184); - 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; -} -} -} -} -else -{ -uint8_t x_286; -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_286 = !lean_is_exclusive(x_16); -if (x_286 == 0) -{ -lean_object* x_287; lean_object* x_288; uint8_t 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; uint8_t x_296; lean_object* x_297; lean_object* x_298; -x_287 = lean_ctor_get(x_16, 0); -lean_dec(x_287); -x_288 = lean_ctor_get(x_7, 5); -lean_inc(x_288); -lean_dec(x_7); -x_289 = 0; -x_290 = l_Lean_SourceInfo_fromRef(x_288, x_289); -x_291 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__6; -lean_inc(x_290); -x_292 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_292, 0, x_290); -lean_ctor_set(x_292, 1, x_291); -x_293 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__5; -x_294 = l_Lean_Syntax_node1(x_290, x_293, x_292); -x_295 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_295, 0, x_294); -x_296 = 1; -x_297 = lean_box(x_296); -x_298 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_298, 0, x_297); -lean_ctor_set(x_298, 1, x_295); -lean_ctor_set(x_16, 0, x_298); -return x_16; -} -else -{ -lean_object* x_299; lean_object* x_300; uint8_t 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; uint8_t x_308; lean_object* x_309; lean_object* x_310; lean_object* x_311; -x_299 = lean_ctor_get(x_16, 1); -lean_inc(x_299); -lean_dec(x_16); -x_300 = lean_ctor_get(x_7, 5); -lean_inc(x_300); -lean_dec(x_7); -x_301 = 0; -x_302 = l_Lean_SourceInfo_fromRef(x_300, x_301); -x_303 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__6; -lean_inc(x_302); -x_304 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_304, 0, x_302); -lean_ctor_set(x_304, 1, x_303); -x_305 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__5; -x_306 = l_Lean_Syntax_node1(x_302, x_305, x_304); -x_307 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_307, 0, x_306); -x_308 = 1; -x_309 = lean_box(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_307); -x_311 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_311, 0, x_310); -lean_ctor_set(x_311, 1, x_299); -return x_311; -} -} -} -else -{ -uint8_t x_312; -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_312 = !lean_is_exclusive(x_16); -if (x_312 == 0) -{ -return x_16; -} -else -{ -lean_object* x_313; lean_object* x_314; lean_object* x_315; -x_313 = lean_ctor_get(x_16, 0); -x_314 = lean_ctor_get(x_16, 1); -lean_inc(x_314); -lean_inc(x_313); -lean_dec(x_16); -x_315 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_315, 0, x_313); -lean_ctor_set(x_315, 1, x_314); -return x_315; -} -} -} -else -{ -uint8_t x_316; -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_316 = !lean_is_exclusive(x_11); -if (x_316 == 0) -{ -lean_object* x_317; lean_object* x_318; -x_317 = lean_ctor_get(x_11, 0); -lean_dec(x_317); -x_318 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -lean_ctor_set(x_11, 0, x_318); -return x_11; -} -else -{ -lean_object* x_319; lean_object* x_320; lean_object* x_321; -x_319 = lean_ctor_get(x_11, 1); -lean_inc(x_319); -lean_dec(x_11); -x_320 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7; -x_321 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_321, 0, x_320); -lean_ctor_set(x_321, 1, x_319); -return x_321; -} -} -} -else -{ -uint8_t x_322; -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_322 = !lean_is_exclusive(x_11); -if (x_322 == 0) -{ -return x_11; -} -else -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; -x_323 = lean_ctor_get(x_11, 0); -x_324 = lean_ctor_get(x_11, 1); -lean_inc(x_324); -lean_inc(x_323); -lean_dec(x_11); -x_325 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_325, 0, x_323); -lean_ctor_set(x_325, 1, x_324); -return x_325; -} -} -} -} -static lean_object* _init_l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabFailureId; -return x_1; -} -} -LEAN_EXPORT lean_object* l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -if (lean_obj_tag(x_3) == 0) -{ -lean_object* x_11; -lean_dec(x_2); -lean_dec(x_1); -x_11 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_10); -return x_11; -} -else -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_12 = lean_ctor_get(x_3, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_3, 1); -lean_inc(x_13); -lean_dec(x_3); -x_14 = lean_box(0); -lean_inc(x_2); -lean_inc(x_1); -x_15 = lean_apply_3(x_12, x_1, x_2, x_14); -if (lean_obj_tag(x_15) == 0) -{ -uint8_t x_16; -lean_dec(x_13); -lean_dec(x_2); -lean_dec(x_1); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; -x_17 = lean_ctor_get(x_15, 1); -lean_dec(x_17); -lean_ctor_set(x_15, 1, x_10); -return x_15; -} -else -{ -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_15, 0); -lean_inc(x_18); -lean_dec(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_10); -return x_19; -} -} -else -{ -lean_object* x_20; uint8_t x_21; -lean_dec(x_15); -x_20 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_10); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_20, 1); -x_24 = l_Lean_Exception_isRuntime(x_22); -if (x_24 == 0) -{ -lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_25 = lean_ctor_get(x_22, 0); -lean_inc(x_25); -x_26 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_27 = lean_nat_dec_eq(x_26, x_25); -lean_dec(x_25); -if (x_27 == 0) -{ -lean_dec(x_13); -lean_dec(x_2); -lean_dec(x_1); -return x_20; -} -else -{ -lean_free_object(x_20); -lean_dec(x_22); -x_3 = x_13; -x_10 = x_23; -goto _start; -} -} -else -{ -uint8_t x_29; -x_29 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -if (x_29 == 0) -{ -lean_dec(x_13); -lean_dec(x_2); -lean_dec(x_1); -return x_20; -} -else -{ -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_22, 0); -lean_inc(x_30); -x_31 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_32 = lean_nat_dec_eq(x_31, x_30); -lean_dec(x_30); -if (x_32 == 0) -{ -lean_dec(x_13); -lean_dec(x_2); -lean_dec(x_1); -return x_20; -} -else -{ -lean_free_object(x_20); -lean_dec(x_22); -x_3 = x_13; -x_10 = x_23; -goto _start; +lean_free_object(x_20); +lean_dec(x_22); +x_3 = x_13; +x_10 = x_23; +goto _start; } } } @@ -15438,576 +15112,955 @@ lean_dec(x_4); return x_11; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___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_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_instInhabitedBool; -x_2 = l_instInhabitedNat; -x_3 = lean_box(x_1); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { -_start: +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -uint8_t x_13; lean_object* x_14; lean_object* x_82; uint8_t x_83; -x_82 = lean_array_get_size(x_3); -x_83 = lean_nat_dec_lt(x_5, x_82); -lean_dec(x_82); -if (x_83 == 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); +x_8 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_7); +lean_dec(x_7); +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; -x_84 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_85 = l___private_Init_GetElem_0__outOfBounds___rarg(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_unbox(x_86); -lean_dec(x_86); -x_13 = x_88; -x_14 = x_87; -goto block_81; +x_3 = x_10; +goto _start; } else { -lean_object* x_89; lean_object* x_90; lean_object* x_91; uint8_t x_92; -x_89 = lean_array_fget(x_3, x_5); -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_unbox(x_90); -lean_dec(x_90); -x_13 = x_92; -x_14 = x_91; -goto block_81; +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_8, 0); +lean_inc(x_12); +lean_dec(x_8); +x_13 = lean_array_push(x_5, x_12); +x_3 = x_10; +x_5 = x_13; +goto _start; +} } -block_81: +else { -lean_object* x_15; uint8_t x_16; -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_nat_dec_eq(x_5, x_15); -if (x_16 == 0) +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_unsigned_to_nat(1u); -x_18 = lean_nat_sub(x_5, x_17); -if (x_13 == 0) +uint8_t x_5; +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_5); -x_39 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_12); -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_19 = x_40; -x_20 = x_41; -goto block_38; +lean_object* x_6; +lean_dec(x_4); +lean_dec(x_3); +x_6 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_6; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = l_Array_extract___rarg(x_2, x_15, x_5); -lean_inc(x_1); -x_43 = l_Lean_Syntax_mkApp(x_1, x_42); -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_4); -x_44 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand(x_4, x_43, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_44) == 0) +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_le(x_4, x_7); +lean_dec(x_7); +if (x_8 == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -lean_dec(x_18); -lean_dec(x_14); +lean_object* x_9; lean_dec(x_4); lean_dec(x_3); -lean_dec(x_1); -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_PrettyPrinter_Delaborator_annotateTermInfo(x_45, x_6, x_7, x_8, x_9, x_10, x_11, 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); -x_48 = !lean_is_exclusive(x_47); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_49 = lean_ctor_get(x_47, 0); -x_50 = lean_array_get_size(x_2); -x_51 = l_Array_extract___rarg(x_2, x_5, x_50); -lean_dec(x_50); -lean_dec(x_2); -x_52 = l_Lean_Syntax_mkApp(x_49, x_51); -lean_ctor_set(x_47, 0, x_52); -return x_47; +x_9 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_9; } 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; -x_53 = lean_ctor_get(x_47, 0); -x_54 = lean_ctor_get(x_47, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_47); -x_55 = lean_array_get_size(x_2); -x_56 = l_Array_extract___rarg(x_2, x_5, x_55); -lean_dec(x_55); -lean_dec(x_2); -x_57 = l_Lean_Syntax_mkApp(x_53, 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_54); -return x_58; +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__2(x_1, x_2, x_10, x_11, x_12); +return x_13; } } -else -{ -lean_object* x_59; lean_object* x_60; -lean_dec(x_5); -x_59 = lean_ctor_get(x_44, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_44, 1); -lean_inc(x_60); -lean_dec(x_44); -x_19 = x_59; -x_20 = x_60; -goto block_38; } } -block_38: -{ -uint8_t x_21; -x_21 = l_Lean_Exception_isRuntime(x_19); -if (x_21 == 0) +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +_start: { -if (lean_obj_tag(x_19) == 0) +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_3, x_4); +if (x_6 == 0) { -lean_object* x_22; -lean_dec(x_18); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); +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); +x_8 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_7); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_19); -lean_ctor_set(x_22, 1, x_20); -return x_22; +x_9 = 1; +x_10 = lean_usize_add(x_3, x_9); +if (lean_obj_tag(x_8) == 0) +{ +x_3 = x_10; +goto _start; } else { -lean_object* x_23; lean_object* x_24; uint8_t x_25; -x_23 = lean_ctor_get(x_19, 0); -lean_inc(x_23); -x_24 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_25 = lean_nat_dec_eq(x_24, x_23); -lean_dec(x_23); -if (x_25 == 0) -{ -lean_object* x_26; -lean_dec(x_18); -lean_dec(x_14); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_8, 0); +lean_inc(x_12); lean_dec(x_8); +x_13 = lean_array_push(x_5, x_12); +x_3 = x_10; +x_5 = x_13; +goto _start; +} +} +else +{ +return x_5; +} +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; +x_5 = lean_nat_dec_lt(x_3, x_4); +if (x_5 == 0) +{ +lean_object* x_6; +lean_dec(x_4); +lean_dec(x_3); +x_6 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_6; +} +else +{ +lean_object* x_7; uint8_t x_8; +x_7 = lean_array_get_size(x_2); +x_8 = lean_nat_dec_le(x_4, x_7); lean_dec(x_7); -lean_dec(x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_19); -lean_ctor_set(x_26, 1, x_20); -return x_26; +x_9 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_9; } else { -lean_object* x_27; lean_object* x_28; -lean_dec(x_19); -x_27 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 12, 5); -lean_closure_set(x_27, 0, x_1); -lean_closure_set(x_27, 1, x_2); -lean_closure_set(x_27, 2, x_3); -lean_closure_set(x_27, 3, x_4); -lean_closure_set(x_27, 4, x_18); -x_28 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_14, x_27, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -return x_28; +size_t x_10; size_t x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_13 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4(x_1, x_2, x_10, x_11, x_12); +return x_13; +} +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("term", 4); +return x_1; +} } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___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_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; uint8_t x_15; +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_4, x_14); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_63; +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_4, x_16); +x_18 = lean_array_get_size(x_2); +x_63 = lean_nat_dec_lt(x_17, x_18); +if (x_63 == 0) +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg; +x_65 = l___private_Init_GetElem_0__outOfBounds___rarg(x_64); +x_66 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_65); +lean_dec(x_65); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; +lean_dec(x_18); +lean_dec(x_4); +x_67 = lean_nat_add(x_5, x_16); +lean_dec(x_5); +x_4 = x_17; +x_5 = x_67; +goto _start; +} +else +{ +lean_object* x_69; +lean_dec(x_66); +x_69 = lean_box(0); +x_19 = x_69; +goto block_62; } } else { -uint8_t x_29; -x_29 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -if (x_29 == 0) +lean_object* x_70; lean_object* x_71; +x_70 = lean_array_fget(x_2, x_17); +x_71 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_70); +lean_dec(x_70); +if (lean_obj_tag(x_71) == 0) { -lean_object* x_30; +lean_object* x_72; lean_dec(x_18); -lean_dec(x_14); +lean_dec(x_4); +x_72 = lean_nat_add(x_5, x_16); +lean_dec(x_5); +x_4 = x_17; +x_5 = x_72; +goto _start; +} +else +{ +lean_object* x_74; +lean_dec(x_71); +x_74 = lean_box(0); +x_19 = x_74; +goto block_62; +} +} +block_62: +{ +lean_object* x_20; lean_object* x_21; lean_object* x_44; lean_object* x_45; +lean_dec(x_19); +lean_inc(x_6); +lean_inc(x_1); +x_44 = l_Lean_Syntax_mkApp(x_1, x_6); +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_3); +x_45 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand(x_3, x_44, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_45) == 0) +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; +lean_dec(x_17); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_1); +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_Array_extract___rarg(x_2, x_4, x_18); +lean_dec(x_18); +lean_dec(x_2); +x_49 = lean_array_get_size(x_48); +x_50 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +x_51 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3(x_50, x_48, x_14, x_49); +lean_dec(x_48); +x_52 = l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(x_46, x_7, x_8, x_9, x_10, x_11, x_12, x_47); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_19); -lean_ctor_set(x_30, 1, x_20); -return x_30; +x_53 = !lean_is_exclusive(x_52); +if (x_53 == 0) +{ +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_52, 0); +x_55 = l_Lean_Syntax_mkApp(x_54, x_51); +lean_ctor_set(x_52, 0, x_55); +return x_52; } else { -if (lean_obj_tag(x_19) == 0) +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_52, 0); +x_57 = lean_ctor_get(x_52, 1); +lean_inc(x_57); +lean_inc(x_56); +lean_dec(x_52); +x_58 = l_Lean_Syntax_mkApp(x_56, x_51); +x_59 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_57); +return x_59; +} +} +else { -lean_object* x_31; +lean_object* x_60; lean_object* x_61; lean_dec(x_18); -lean_dec(x_14); +lean_dec(x_4); +x_60 = lean_ctor_get(x_45, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_45, 1); +lean_inc(x_61); +lean_dec(x_45); +x_20 = x_60; +x_21 = x_61; +goto block_43; +} +block_43: +{ +uint8_t x_22; +x_22 = l_Lean_Exception_isRuntime(x_20); +if (x_22 == 0) +{ +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_23; +lean_dec(x_17); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_19); -lean_ctor_set(x_31, 1, x_20); -return x_31; +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_21); +return x_23; } else { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_ctor_get(x_19, 0); -lean_inc(x_32); -x_33 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_34 = lean_nat_dec_eq(x_33, x_32); -lean_dec(x_32); -if (x_34 == 0) +lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_24 = lean_ctor_get(x_20, 0); +lean_inc(x_24); +x_25 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_26 = lean_nat_dec_eq(x_25, x_24); +lean_dec(x_24); +if (x_26 == 0) { -lean_object* x_35; -lean_dec(x_18); -lean_dec(x_14); +lean_object* x_27; +lean_dec(x_17); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_19); -lean_ctor_set(x_35, 1, x_20); -return x_35; +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_21); +return x_27; } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_19); -x_36 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 12, 5); -lean_closure_set(x_36, 0, x_1); -lean_closure_set(x_36, 1, x_2); -lean_closure_set(x_36, 2, x_3); -lean_closure_set(x_36, 3, x_4); -lean_closure_set(x_36, 4, x_18); -x_37 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_14, x_36, x_6, x_7, x_8, x_9, x_10, x_11, x_20); -return x_37; -} -} -} +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_20); +x_28 = lean_nat_add(x_5, x_16); +lean_dec(x_5); +x_29 = lean_array_pop(x_6); +x_30 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 13, 6); +lean_closure_set(x_30, 0, x_1); +lean_closure_set(x_30, 1, x_2); +lean_closure_set(x_30, 2, x_3); +lean_closure_set(x_30, 3, x_17); +lean_closure_set(x_30, 4, x_14); +lean_closure_set(x_30, 5, x_29); +x_31 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_28, x_30, x_7, x_8, x_9, x_10, x_11, x_12, x_21); +return x_31; } } } else { -lean_dec(x_14); +uint8_t x_32; +x_32 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); +if (x_32 == 0) +{ +lean_object* x_33; +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -if (x_13 == 0) +lean_dec(x_2); +lean_dec(x_1); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_21); +return x_33; +} +else { -lean_object* x_61; uint8_t x_62; +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_34; +lean_dec(x_17); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); +lean_dec(x_5); +lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_61 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_12); -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_20); +lean_ctor_set(x_34, 1, x_21); +return x_34; +} +else { -return x_61; +lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_35 = lean_ctor_get(x_20, 0); +lean_inc(x_35); +x_36 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_37 = lean_nat_dec_eq(x_36, x_35); +lean_dec(x_35); +if (x_37 == 0) +{ +lean_object* x_38; +lean_dec(x_17); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_20); +lean_ctor_set(x_38, 1, x_21); +return x_38; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_61, 0); -x_64 = lean_ctor_get(x_61, 1); -lean_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; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_dec(x_20); +x_39 = lean_nat_add(x_5, x_16); +lean_dec(x_5); +x_40 = lean_array_pop(x_6); +x_41 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 13, 6); +lean_closure_set(x_41, 0, x_1); +lean_closure_set(x_41, 1, x_2); +lean_closure_set(x_41, 2, x_3); +lean_closure_set(x_41, 3, x_17); +lean_closure_set(x_41, 4, x_14); +lean_closure_set(x_41, 5, x_40); +x_42 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_39, x_41, x_7, x_8, x_9, x_10, x_11, x_12, x_21); +return x_42; +} +} +} +} +} } } else { -lean_object* x_66; +lean_object* x_75; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +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_66 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand(x_4, x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_66) == 0) +x_75 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand(x_3, x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_75) == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; uint8_t x_70; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(x_67, x_6, x_7, x_8, x_9, x_10, x_11, x_68); +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_78 = l_Lean_PrettyPrinter_Delaborator_annotateTermInfo(x_76, x_7, x_8, x_9, x_10, x_11, x_12, x_77); +lean_dec(x_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_70 = !lean_is_exclusive(x_69); -if (x_70 == 0) +x_79 = !lean_is_exclusive(x_78); +if (x_79 == 0) { -lean_object* x_71; lean_object* x_72; -x_71 = lean_ctor_get(x_69, 0); -x_72 = l_Lean_Syntax_mkApp(x_71, x_2); -lean_ctor_set(x_69, 0, x_72); -return x_69; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_80 = lean_ctor_get(x_78, 0); +x_81 = lean_array_get_size(x_2); +x_82 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +x_83 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(x_82, x_2, x_14, x_81); +lean_dec(x_2); +x_84 = l_Lean_Syntax_mkApp(x_80, x_83); +lean_ctor_set(x_78, 0, x_84); +return x_78; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_73 = lean_ctor_get(x_69, 0); -x_74 = lean_ctor_get(x_69, 1); -lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_69); -x_75 = l_Lean_Syntax_mkApp(x_73, x_2); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -return x_76; +lean_object* x_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_85 = lean_ctor_get(x_78, 0); +x_86 = lean_ctor_get(x_78, 1); +lean_inc(x_86); +lean_inc(x_85); +lean_dec(x_78); +x_87 = lean_array_get_size(x_2); +x_88 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +x_89 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(x_88, x_2, x_14, x_87); +lean_dec(x_2); +x_90 = l_Lean_Syntax_mkApp(x_85, x_89); +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_90); +lean_ctor_set(x_91, 1, x_86); +return x_91; } } else { -uint8_t x_77; +uint8_t x_92; +lean_dec(x_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_77 = !lean_is_exclusive(x_66); -if (x_77 == 0) +x_92 = !lean_is_exclusive(x_75); +if (x_92 == 0) { -return x_66; +return x_75; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_66, 0); -x_79 = lean_ctor_get(x_66, 1); -lean_inc(x_79); -lean_inc(x_78); -lean_dec(x_66); -x_80 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -return x_80; -} -} +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_75, 0); +x_94 = lean_ctor_get(x_75, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_75); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_13; lean_object* x_14; -x_13 = lean_array_get_size(x_1); -x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go(x_2, x_1, x_3, x_4, x_13, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_14; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__2(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_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1() { +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1___boxed(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("Lean.PrettyPrinter.Delaborator.delabAppImplicitCore.tryAppUnexpanders", 69); -return x_1; +lean_object* x_5; +x_5 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___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_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_delabAppImplicitCore_tryAppUnexpanders___closed__1; -x_3 = lean_unsigned_to_nat(418u); -x_4 = lean_unsigned_to_nat(58u); -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; +size_t x_6; size_t x_7; lean_object* x_8; +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__4(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_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3() { +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_1; -x_1 = l_Lean_PrettyPrinter_Delaborator_appUnexpanderAttribute; -return x_1; +lean_object* x_5; +x_5 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_2); +lean_dec(x_1); +return x_5; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_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_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _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_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_14 = l_Lean_Expr_getAppFn(x_12); -lean_dec(x_12); -x_15 = l_Lean_Expr_consumeMData(x_14); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 4) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 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; uint8_t x_23; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -lean_dec(x_15); -x_17 = lean_st_ref_get(x_9, 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 = lean_ctor_get(x_18, 0); -lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3; -x_22 = l_Lean_KeyedDeclsAttribute_getValues___rarg(x_21, x_20, x_16); -x_23 = l_List_isEmpty___rarg(x_22); -if (x_23 == 0) +lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_6); +lean_dec(x_6); +x_8 = 1; +x_9 = lean_usize_add(x_2, x_8); +if (lean_obj_tag(x_7) == 0) { -lean_object* x_24; lean_object* x_25; -x_24 = lean_box(0); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(x_2, x_1, x_3, x_22, x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -return x_25; +x_2 = x_9; +goto _start; } else { -lean_object* x_26; uint8_t x_27; -lean_dec(x_22); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); 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_26 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_19); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -return x_26; +x_12 = lean_array_push(x_4, x_11); +x_2 = x_9; +x_4 = x_12; +goto _start; +} } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +return x_4; } } } +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +lean_dec(x_2); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_5; +} else { -lean_object* x_31; lean_object* x_32; -lean_dec(x_15); +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_le(x_3, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_31 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2; -x_32 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_32; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_8; +} +else +{ +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__2(x_1, x_9, x_10, x_11); +return x_12; +} } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___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_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(lean_object* x_1) { _start: { -lean_object* x_13; -x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_5); -return x_13; +uint8_t x_2; +x_2 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand(x_1); +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; } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_12 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1; +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Array_findIdx_x3f_loop___rarg(x_1, x_12, x_13); +x_15 = lean_array_get_size(x_1); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_16 = lean_nat_sub(x_15, x_15); +x_17 = l_Array_extract___rarg(x_1, x_13, x_15); +x_18 = lean_array_get_size(x_17); +x_19 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1(x_17, x_13, x_18); +lean_dec(x_17); +x_20 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 13, 6); +lean_closure_set(x_20, 0, x_2); +lean_closure_set(x_20, 1, x_1); +lean_closure_set(x_20, 2, x_3); +lean_closure_set(x_20, 3, x_15); +lean_closure_set(x_20, 4, x_13); +lean_closure_set(x_20, 5, x_19); +x_21 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_16, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_21; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_22 = lean_ctor_get(x_14, 0); +lean_inc(x_22); +lean_dec(x_14); +x_23 = lean_nat_sub(x_15, x_22); +lean_dec(x_15); +x_24 = l_Array_extract___rarg(x_1, x_13, x_22); +x_25 = lean_array_get_size(x_24); +x_26 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1(x_24, x_13, x_25); +lean_dec(x_24); +x_27 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go), 13, 6); +lean_closure_set(x_27, 0, x_2); +lean_closure_set(x_27, 1, x_1); +lean_closure_set(x_27, 2, x_3); +lean_closure_set(x_27, 3, x_22); +lean_closure_set(x_27, 4, x_13); +lean_closure_set(x_27, 5, x_26); +x_28 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_23, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_28; +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.PrettyPrinter.Delaborator.delabAppImplicitCore.tryAppUnexpanders", 69); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1; +x_3 = lean_unsigned_to_nat(453u); +x_4 = lean_unsigned_to_nat(58u); +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_delabAppImplicitCore_tryAppUnexpanders___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_PrettyPrinter_Delaborator_appUnexpanderAttribute; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Expr_getAppFn(x_11); +lean_dec(x_11); +x_14 = l_Lean_Expr_consumeMData(x_13); +lean_dec(x_13); +if (lean_obj_tag(x_14) == 4) +{ +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_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_st_ref_get(x_8, 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 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3; +x_21 = l_Lean_KeyedDeclsAttribute_getValues___rarg(x_20, x_19, x_15); +x_22 = l_List_isEmpty___rarg(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +x_23 = lean_box(0); +x_24 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2(x_2, x_1, x_21, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_24; +} +else +{ +lean_object* x_25; uint8_t x_26; +lean_dec(x_21); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_25 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_18); +x_26 = !lean_is_exclusive(x_25); +if (x_26 == 0) +{ +return x_25; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_25, 0); +x_28 = lean_ctor_get(x_25, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_25); +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; +} +} +} +else +{ +lean_object* x_30; lean_object* x_31; +lean_dec(x_14); +lean_dec(x_2); +lean_dec(x_1); +x_30 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2; +x_31 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1(x_30, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___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_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__2(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___spec__1(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__1(x_1); +lean_dec(x_1); +x_3 = lean_box(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___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_delabAppImplicitCore_tryAppUnexpanders___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_4); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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; @@ -16192,1797 +16245,1302 @@ return x_29; } } } -LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(lean_object* x_1, size_t x_2, size_t x_3) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -uint8_t x_4; -x_4 = lean_usize_dec_eq(x_2, x_3); -if (x_4 == 0) +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_array_uget(x_1, x_2); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_unbox(x_6); +lean_object* x_6; lean_object* x_7; size_t x_8; size_t x_9; +x_6 = lean_array_uget(x_1, x_2); +x_7 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_syntax_x3f(x_6); 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_2, x_8); +if (lean_obj_tag(x_7) == 0) +{ x_2 = x_9; goto _start; } else { -uint8_t x_11; -x_11 = 1; -return x_11; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_7, 0); +lean_inc(x_11); +lean_dec(x_7); +x_12 = lean_array_push(x_4, x_11); +x_2 = x_9; +x_4 = x_12; +goto _start; } } else { -uint8_t x_12; -x_12 = 0; -return x_12; +return x_4; } } } -static lean_object* _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1() { +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_instInhabitedNat; -x_2 = l_Lean_instInhabitedName; -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_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(lean_object* x_1) { -_start: +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) { -lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; +lean_object* x_5; +lean_dec(x_3); +lean_dec(x_2); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_5; } +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_le(x_3, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_3); +lean_dec(x_2); +x_8 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_8; } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_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; lean_object* x_26; -x_12 = lean_box(0); -x_13 = lean_array_to_list(lean_box(0), x_1); -x_14 = lean_mk_empty_array_with_capacity(x_2); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_box(x_3); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_15); -lean_inc(x_14); -x_18 = lean_array_push(x_14, x_17); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_14); -lean_ctor_set(x_19, 1, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_13); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_15); -lean_ctor_set(x_21, 1, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_12); -lean_ctor_set(x_22, 1, x_21); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_4); -lean_ctor_set(x_23, 1, x_22); -x_24 = lean_box(x_3); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_23); -x_26 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_26, 0, x_25); -lean_ctor_set(x_26, 1, x_11); -return x_26; +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(x_1, x_9, x_10, x_11); +return x_12; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = lean_unsigned_to_nat(1u); -x_3 = lean_box(x_1); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7(lean_object* x_1, size_t x_2, size_t x_3) { _start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_nat_sub(x_1, x_2); -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_17, x_18); -lean_dec(x_17); -lean_inc(x_3); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx(x_19, x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); if (x_4 == 0) { -uint8_t x_22; -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, 1); -x_24 = lean_ctor_get(x_21, 0); -lean_dec(x_24); -if (lean_obj_tag(x_23) == 0) -{ -uint8_t x_25; -x_25 = !lean_is_exclusive(x_20); -if (x_25 == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_26 = lean_ctor_get(x_20, 0); -lean_dec(x_26); -x_27 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_28 = l_Array_back___rarg(x_27, x_5); -x_29 = !lean_is_exclusive(x_28); -if (x_29 == 0) +lean_object* x_5; +x_5 = lean_array_uget(x_1, x_2); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_30 = lean_ctor_get(x_28, 1); -x_31 = lean_ctor_get(x_28, 0); -lean_dec(x_31); -x_32 = lean_array_pop(x_5); -x_33 = lean_nat_add(x_30, x_18); -lean_dec(x_30); -x_34 = 0; -x_35 = lean_box(x_34); -lean_ctor_set(x_28, 1, x_33); -lean_ctor_set(x_28, 0, x_35); -x_36 = lean_array_push(x_32, x_28); -x_37 = lean_nat_add(x_2, x_18); -x_38 = lean_box(0); -x_39 = l_List_tailD___rarg(x_3, x_38); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_36); -lean_ctor_set(x_21, 0, x_6); -x_40 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_21); -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_7); -lean_ctor_set(x_42, 1, x_41); -x_43 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_43, 0, x_8); -lean_ctor_set(x_43, 1, x_42); -x_44 = lean_box(x_34); -x_45 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_45, 0, x_44); -lean_ctor_set(x_45, 1, x_43); -lean_ctor_set(x_20, 0, x_45); -return x_20; +size_t x_6; size_t x_7; +x_6 = 1; +x_7 = lean_usize_add(x_2, x_6); +x_2 = x_7; +goto _start; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; 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_46 = lean_ctor_get(x_28, 1); -lean_inc(x_46); -lean_dec(x_28); -x_47 = lean_array_pop(x_5); -x_48 = lean_nat_add(x_46, x_18); -lean_dec(x_46); -x_49 = 0; -x_50 = lean_box(x_49); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_48); -x_52 = lean_array_push(x_47, x_51); -x_53 = lean_nat_add(x_2, x_18); -x_54 = lean_box(0); -x_55 = l_List_tailD___rarg(x_3, x_54); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_52); -lean_ctor_set(x_21, 0, x_6); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_21); -x_57 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_57, 0, x_53); -lean_ctor_set(x_57, 1, x_56); -x_58 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_58, 0, x_7); -lean_ctor_set(x_58, 1, x_57); -x_59 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_59, 0, x_8); -lean_ctor_set(x_59, 1, x_58); -x_60 = lean_box(x_49); -x_61 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_61, 0, x_60); -lean_ctor_set(x_61, 1, x_59); -lean_ctor_set(x_20, 0, x_61); -return x_20; +uint8_t x_9; +lean_dec(x_5); +x_9 = 1; +return x_9; } } 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; 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; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -x_62 = lean_ctor_get(x_20, 1); -lean_inc(x_62); -lean_dec(x_20); -x_63 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_64 = l_Array_back___rarg(x_63, x_5); -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -if (lean_is_exclusive(x_64)) { - lean_ctor_release(x_64, 0); - lean_ctor_release(x_64, 1); - x_66 = x_64; -} else { - lean_dec_ref(x_64); - x_66 = lean_box(0); -} -x_67 = lean_array_pop(x_5); -x_68 = lean_nat_add(x_65, x_18); -lean_dec(x_65); -x_69 = 0; -x_70 = lean_box(x_69); -if (lean_is_scalar(x_66)) { - x_71 = lean_alloc_ctor(0, 2, 0); -} else { - x_71 = x_66; +uint8_t x_10; +x_10 = 0; +return x_10; } -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_68); -x_72 = lean_array_push(x_67, x_71); -x_73 = lean_nat_add(x_2, x_18); -x_74 = lean_box(0); -x_75 = l_List_tailD___rarg(x_3, x_74); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_72); -lean_ctor_set(x_21, 0, x_6); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_21); -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_73); -lean_ctor_set(x_77, 1, x_76); -x_78 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_78, 0, x_7); -lean_ctor_set(x_78, 1, x_77); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_8); -lean_ctor_set(x_79, 1, x_78); -x_80 = lean_box(x_69); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, 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_62); -return x_82; } } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { +_start: { -uint8_t x_83; -x_83 = !lean_is_exclusive(x_20); -if (x_83 == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_11, x_13); +x_15 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1; +lean_inc(x_14); +x_16 = lean_mk_array(x_14, x_15); +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_14, x_17); +lean_dec(x_14); +x_19 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_11, x_16, x_18); +x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_getOptionsAtCurrPos___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_array_get_size(x_19); +x_24 = l_Lean_SubExpr_Pos_pushNaryArg(x_23, x_1, x_21); +lean_dec(x_21); +x_25 = !lean_is_exclusive(x_3); +if (x_25 == 0) { -lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; uint8_t x_96; lean_object* x_97; lean_object* x_98; -x_84 = lean_ctor_get(x_20, 0); -lean_dec(x_84); -x_85 = lean_ctor_get(x_23, 0); -lean_inc(x_85); +lean_object* x_26; uint8_t x_27; +x_26 = lean_ctor_get(x_3, 3); +lean_dec(x_26); +x_27 = lean_nat_dec_lt(x_1, x_23); lean_dec(x_23); -x_86 = lean_array_push(x_6, x_85); -x_87 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; -x_88 = lean_array_push(x_5, x_87); -x_89 = lean_nat_add(x_2, x_18); -x_90 = lean_box(0); -x_91 = l_List_tailD___rarg(x_3, x_90); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_88); -lean_ctor_set(x_21, 0, x_86); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_91); -lean_ctor_set(x_92, 1, x_21); -x_93 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_93, 0, x_89); -lean_ctor_set(x_93, 1, x_92); -x_94 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_94, 0, x_7); -lean_ctor_set(x_94, 1, x_93); -x_95 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_95, 0, x_8); -lean_ctor_set(x_95, 1, x_94); -x_96 = 0; -x_97 = lean_box(x_96); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_95); -lean_ctor_set(x_20, 0, x_98); -return x_20; +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +lean_dec(x_19); +x_28 = l_Lean_instInhabitedExpr; +x_29 = l___private_Init_GetElem_0__outOfBounds___rarg(x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_24); +lean_ctor_set(x_3, 3, x_30); +x_31 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +return x_31; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; -x_99 = lean_ctor_get(x_20, 1); -lean_inc(x_99); -lean_dec(x_20); -x_100 = lean_ctor_get(x_23, 0); -lean_inc(x_100); -lean_dec(x_23); -x_101 = lean_array_push(x_6, x_100); -x_102 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; -x_103 = lean_array_push(x_5, x_102); -x_104 = lean_nat_add(x_2, x_18); -x_105 = lean_box(0); -x_106 = l_List_tailD___rarg(x_3, x_105); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_103); -lean_ctor_set(x_21, 0, x_101); -x_107 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_107, 0, x_106); -lean_ctor_set(x_107, 1, x_21); -x_108 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_108, 0, x_104); -lean_ctor_set(x_108, 1, x_107); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_7); -lean_ctor_set(x_109, 1, x_108); -x_110 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_110, 0, x_8); -lean_ctor_set(x_110, 1, x_109); -x_111 = 0; -x_112 = lean_box(x_111); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_112); -lean_ctor_set(x_113, 1, x_110); -x_114 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_114, 0, x_113); -lean_ctor_set(x_114, 1, x_99); -return x_114; -} +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_array_fget(x_19, x_1); +lean_dec(x_19); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_24); +lean_ctor_set(x_3, 3, x_33); +x_34 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +return x_34; } } else { -lean_object* x_115; -x_115 = lean_ctor_get(x_21, 1); -lean_inc(x_115); -lean_dec(x_21); -if (lean_obj_tag(x_115) == 0) +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; uint8_t x_40; +x_35 = lean_ctor_get(x_3, 0); +x_36 = lean_ctor_get(x_3, 1); +x_37 = lean_ctor_get(x_3, 2); +x_38 = lean_ctor_get_uint8(x_3, sizeof(void*)*5); +x_39 = lean_ctor_get(x_3, 4); +lean_inc(x_39); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_3); +x_40 = lean_nat_dec_lt(x_1, x_23); +lean_dec(x_23); +if (x_40 == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t 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_116 = lean_ctor_get(x_20, 1); -lean_inc(x_116); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - lean_ctor_release(x_20, 1); - x_117 = x_20; -} else { - lean_dec_ref(x_20); - x_117 = lean_box(0); -} -x_118 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_119 = l_Array_back___rarg(x_118, x_5); -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_array_pop(x_5); -x_123 = lean_nat_add(x_120, x_18); -lean_dec(x_120); -x_124 = 0; -x_125 = lean_box(x_124); -if (lean_is_scalar(x_121)) { - x_126 = lean_alloc_ctor(0, 2, 0); -} else { - x_126 = x_121; -} -lean_ctor_set(x_126, 0, x_125); -lean_ctor_set(x_126, 1, x_123); -x_127 = lean_array_push(x_122, x_126); -x_128 = lean_nat_add(x_2, x_18); -x_129 = lean_box(0); -x_130 = l_List_tailD___rarg(x_3, x_129); -lean_dec(x_3); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_6); -lean_ctor_set(x_131, 1, x_127); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -x_133 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_133, 0, x_128); -lean_ctor_set(x_133, 1, x_132); -x_134 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_134, 0, x_7); -lean_ctor_set(x_134, 1, x_133); -x_135 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_135, 0, x_8); -lean_ctor_set(x_135, 1, x_134); -x_136 = lean_box(x_124); -x_137 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_137, 0, x_136); -lean_ctor_set(x_137, 1, x_135); -if (lean_is_scalar(x_117)) { - x_138 = lean_alloc_ctor(0, 2, 0); -} else { - x_138 = x_117; -} -lean_ctor_set(x_138, 0, x_137); -lean_ctor_set(x_138, 1, x_116); -return x_138; +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +lean_dec(x_19); +x_41 = l_Lean_instInhabitedExpr; +x_42 = l___private_Init_GetElem_0__outOfBounds___rarg(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_24); +x_44 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_44, 0, x_35); +lean_ctor_set(x_44, 1, x_36); +lean_ctor_set(x_44, 2, x_37); +lean_ctor_set(x_44, 3, x_43); +lean_ctor_set(x_44, 4, x_39); +lean_ctor_set_uint8(x_44, sizeof(void*)*5, x_38); +x_45 = lean_apply_7(x_2, x_44, x_4, x_5, x_6, x_7, x_8, x_22); +return x_45; } 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; uint8_t x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; -x_139 = lean_ctor_get(x_20, 1); -lean_inc(x_139); -if (lean_is_exclusive(x_20)) { - lean_ctor_release(x_20, 0); - lean_ctor_release(x_20, 1); - x_140 = x_20; -} else { - lean_dec_ref(x_20); - x_140 = lean_box(0); -} -x_141 = lean_ctor_get(x_115, 0); -lean_inc(x_141); -lean_dec(x_115); -x_142 = lean_array_push(x_6, x_141); -x_143 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; -x_144 = lean_array_push(x_5, x_143); -x_145 = lean_nat_add(x_2, x_18); -x_146 = lean_box(0); -x_147 = l_List_tailD___rarg(x_3, x_146); -lean_dec(x_3); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_142); -lean_ctor_set(x_148, 1, x_144); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_147); -lean_ctor_set(x_149, 1, x_148); -x_150 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_150, 0, x_145); -lean_ctor_set(x_150, 1, x_149); -x_151 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_151, 0, x_7); -lean_ctor_set(x_151, 1, x_150); -x_152 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_152, 0, x_8); -lean_ctor_set(x_152, 1, x_151); -x_153 = 0; -x_154 = lean_box(x_153); -x_155 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_155, 0, x_154); -lean_ctor_set(x_155, 1, x_152); -if (lean_is_scalar(x_140)) { - x_156 = lean_alloc_ctor(0, 2, 0); -} else { - x_156 = x_140; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_46 = lean_array_fget(x_19, x_1); +lean_dec(x_19); +x_47 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_24); +x_48 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_48, 0, x_35); +lean_ctor_set(x_48, 1, x_36); +lean_ctor_set(x_48, 2, x_37); +lean_ctor_set(x_48, 3, x_47); +lean_ctor_set(x_48, 4, x_39); +lean_ctor_set_uint8(x_48, sizeof(void*)*5, x_38); +x_49 = lean_apply_7(x_2, x_48, x_4, x_5, x_6, x_7, x_8, x_22); +return x_49; } -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_139); -return x_156; } } } -else -{ -lean_object* x_157; -x_157 = lean_ctor_get(x_21, 1); -lean_inc(x_157); -if (lean_obj_tag(x_157) == 0) +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9(lean_object* x_1, size_t x_2, size_t x_3) { +_start: { -uint8_t x_158; -x_158 = !lean_is_exclusive(x_20); -if (x_158 == 0) -{ -lean_object* x_159; uint8_t x_160; -x_159 = lean_ctor_get(x_20, 0); -lean_dec(x_159); -x_160 = !lean_is_exclusive(x_21); -if (x_160 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; uint8_t x_165; -x_161 = lean_ctor_get(x_21, 0); -x_162 = lean_ctor_get(x_21, 1); -lean_dec(x_162); -x_163 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_164 = l_Array_back___rarg(x_163, x_5); -x_165 = !lean_is_exclusive(x_164); -if (x_165 == 0) +lean_object* x_5; uint8_t x_6; +x_5 = lean_array_uget(x_1, x_2); +x_6 = l_Lean_PrettyPrinter_Delaborator_AppImplicitArg_canUnexpand(x_5); +lean_dec(x_5); +if (x_6 == 0) { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_166 = lean_ctor_get(x_164, 1); -x_167 = lean_ctor_get(x_164, 0); -lean_dec(x_167); -x_168 = lean_array_pop(x_5); -x_169 = lean_nat_add(x_166, x_18); -lean_dec(x_166); -lean_inc(x_161); -lean_ctor_set(x_164, 1, x_169); -lean_ctor_set(x_164, 0, x_161); -x_170 = lean_array_push(x_168, x_164); -x_171 = lean_nat_add(x_2, x_18); -x_172 = lean_box(0); -x_173 = l_List_tailD___rarg(x_3, x_172); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_170); -lean_ctor_set(x_21, 0, x_6); -x_174 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_174, 0, x_173); -lean_ctor_set(x_174, 1, x_21); -x_175 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_175, 0, x_171); -lean_ctor_set(x_175, 1, x_174); -x_176 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_176, 0, x_7); -lean_ctor_set(x_176, 1, x_175); -x_177 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_177, 0, x_8); -lean_ctor_set(x_177, 1, x_176); -x_178 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_178, 0, x_161); -lean_ctor_set(x_178, 1, x_177); -lean_ctor_set(x_20, 0, x_178); -return x_20; +uint8_t x_7; +x_7 = 1; +return x_7; } else { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -x_179 = lean_ctor_get(x_164, 1); -lean_inc(x_179); -lean_dec(x_164); -x_180 = lean_array_pop(x_5); -x_181 = lean_nat_add(x_179, x_18); -lean_dec(x_179); -lean_inc(x_161); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_161); -lean_ctor_set(x_182, 1, x_181); -x_183 = lean_array_push(x_180, x_182); -x_184 = lean_nat_add(x_2, x_18); -x_185 = lean_box(0); -x_186 = l_List_tailD___rarg(x_3, x_185); -lean_dec(x_3); -lean_ctor_set(x_21, 1, x_183); -lean_ctor_set(x_21, 0, x_6); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_21); -x_188 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_188, 0, x_184); -lean_ctor_set(x_188, 1, x_187); -x_189 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_189, 0, x_7); -lean_ctor_set(x_189, 1, x_188); -x_190 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_190, 0, x_8); -lean_ctor_set(x_190, 1, x_189); -x_191 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_191, 0, x_161); -lean_ctor_set(x_191, 1, x_190); -lean_ctor_set(x_20, 0, x_191); -return x_20; +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 { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -x_192 = lean_ctor_get(x_21, 0); -lean_inc(x_192); -lean_dec(x_21); -x_193 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_194 = l_Array_back___rarg(x_193, x_5); -x_195 = lean_ctor_get(x_194, 1); -lean_inc(x_195); -if (lean_is_exclusive(x_194)) { - lean_ctor_release(x_194, 0); - lean_ctor_release(x_194, 1); - x_196 = x_194; -} else { - lean_dec_ref(x_194); - x_196 = lean_box(0); +uint8_t x_11; +x_11 = 0; +return x_11; } -x_197 = lean_array_pop(x_5); -x_198 = lean_nat_add(x_195, x_18); -lean_dec(x_195); -lean_inc(x_192); -if (lean_is_scalar(x_196)) { - x_199 = lean_alloc_ctor(0, 2, 0); -} else { - x_199 = x_196; -} -lean_ctor_set(x_199, 0, x_192); -lean_ctor_set(x_199, 1, x_198); -x_200 = lean_array_push(x_197, x_199); -x_201 = lean_nat_add(x_2, x_18); -x_202 = lean_box(0); -x_203 = l_List_tailD___rarg(x_3, x_202); -lean_dec(x_3); -x_204 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_204, 0, x_6); -lean_ctor_set(x_204, 1, x_200); -x_205 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -x_206 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_206, 0, x_201); -lean_ctor_set(x_206, 1, x_205); -x_207 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_207, 0, x_7); -lean_ctor_set(x_207, 1, x_206); -x_208 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_208, 0, x_8); -lean_ctor_set(x_208, 1, x_207); -x_209 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_209, 0, x_192); -lean_ctor_set(x_209, 1, x_208); -lean_ctor_set(x_20, 0, x_209); -return x_20; } } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; -x_210 = lean_ctor_get(x_20, 1); -lean_inc(x_210); -lean_dec(x_20); -x_211 = lean_ctor_get(x_21, 0); -lean_inc(x_211); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_212 = x_21; -} else { - lean_dec_ref(x_21); - x_212 = lean_box(0); -} -x_213 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_214 = l_Array_back___rarg(x_213, x_5); -x_215 = lean_ctor_get(x_214, 1); -lean_inc(x_215); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_216 = x_214; -} else { - lean_dec_ref(x_214); - x_216 = lean_box(0); -} -x_217 = lean_array_pop(x_5); -x_218 = lean_nat_add(x_215, x_18); -lean_dec(x_215); -lean_inc(x_211); -if (lean_is_scalar(x_216)) { - x_219 = lean_alloc_ctor(0, 2, 0); -} else { - x_219 = x_216; -} -lean_ctor_set(x_219, 0, x_211); -lean_ctor_set(x_219, 1, x_218); -x_220 = lean_array_push(x_217, x_219); -x_221 = lean_nat_add(x_2, x_18); -x_222 = lean_box(0); -x_223 = l_List_tailD___rarg(x_3, x_222); -lean_dec(x_3); -if (lean_is_scalar(x_212)) { - x_224 = lean_alloc_ctor(0, 2, 0); -} else { - x_224 = x_212; -} -lean_ctor_set(x_224, 0, x_6); -lean_ctor_set(x_224, 1, x_220); -x_225 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -x_226 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_226, 0, x_221); -lean_ctor_set(x_226, 1, x_225); -x_227 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_227, 0, x_7); -lean_ctor_set(x_227, 1, x_226); -x_228 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_228, 0, x_8); -lean_ctor_set(x_228, 1, x_227); -x_229 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_229, 0, x_211); -lean_ctor_set(x_229, 1, x_228); -x_230 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_210); -return x_230; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_mk_empty_array_with_capacity(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_2); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_9); +return x_12; } } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_231; -x_231 = !lean_is_exclusive(x_20); -if (x_231 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_3); +if (x_11 == 0) { -lean_object* x_232; uint8_t x_233; -x_232 = lean_ctor_get(x_20, 0); -lean_dec(x_232); -x_233 = !lean_is_exclusive(x_21); -if (x_233 == 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_3, 0); +x_13 = lean_ctor_get(x_3, 1); +x_14 = lean_array_get_size(x_13); +x_15 = lean_nat_sub(x_1, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_15, x_16); +lean_dec(x_15); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_dec_lt(x_14, x_18); +lean_dec(x_18); +if (x_19 == 0) { -lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_234 = lean_ctor_get(x_21, 0); -x_235 = lean_ctor_get(x_21, 1); -lean_dec(x_235); -x_236 = lean_ctor_get(x_157, 0); -lean_inc(x_236); -lean_dec(x_157); -x_237 = lean_array_push(x_6, x_236); -lean_inc(x_234); -lean_ctor_set(x_21, 1, x_18); -x_238 = lean_array_push(x_5, x_21); -x_239 = lean_nat_add(x_2, x_18); -x_240 = lean_box(0); -x_241 = l_List_tailD___rarg(x_3, x_240); -lean_dec(x_3); -x_242 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_242, 0, x_237); -lean_ctor_set(x_242, 1, x_238); -x_243 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_243, 0, x_241); -lean_ctor_set(x_243, 1, x_242); -x_244 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_244, 0, x_239); -lean_ctor_set(x_244, 1, x_243); -x_245 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_245, 0, x_7); -lean_ctor_set(x_245, 1, x_244); -x_246 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_246, 0, x_8); -lean_ctor_set(x_246, 1, x_245); -x_247 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_247, 0, x_234); -lean_ctor_set(x_247, 1, x_246); -lean_ctor_set(x_20, 0, x_247); -return x_20; -} -else +lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_dec(x_14); +x_20 = l_Lean_PrettyPrinter_Delaborator_instInhabitedParamKind; +x_21 = l___private_Init_GetElem_0__outOfBounds___rarg(x_20); +x_22 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(x_17, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_17); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_248 = lean_ctor_get(x_21, 0); -lean_inc(x_248); -lean_dec(x_21); -x_249 = lean_ctor_get(x_157, 0); -lean_inc(x_249); -lean_dec(x_157); -x_250 = lean_array_push(x_6, x_249); -lean_inc(x_248); -x_251 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_251, 0, x_248); -lean_ctor_set(x_251, 1, x_18); -x_252 = lean_array_push(x_5, x_251); -x_253 = lean_nat_add(x_2, x_18); -x_254 = lean_box(0); -x_255 = l_List_tailD___rarg(x_3, x_254); -lean_dec(x_3); -x_256 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_256, 0, x_250); -lean_ctor_set(x_256, 1, x_252); -x_257 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_257, 0, x_255); -lean_ctor_set(x_257, 1, x_256); -x_258 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_258, 0, x_253); -lean_ctor_set(x_258, 1, x_257); -x_259 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_259, 0, x_7); -lean_ctor_set(x_259, 1, x_258); -x_260 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_260, 0, x_8); -lean_ctor_set(x_260, 1, x_259); -x_261 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_261, 0, x_248); -lean_ctor_set(x_261, 1, x_260); -lean_ctor_set(x_20, 0, x_261); -return x_20; -} +uint8_t x_23; +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); +x_25 = lean_array_push(x_13, x_24); +lean_ctor_set(x_3, 1, x_25); +lean_ctor_set(x_22, 0, x_3); +return x_22; } else { -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; -x_262 = lean_ctor_get(x_20, 1); -lean_inc(x_262); -lean_dec(x_20); -x_263 = lean_ctor_get(x_21, 0); -lean_inc(x_263); -if (lean_is_exclusive(x_21)) { - lean_ctor_release(x_21, 0); - lean_ctor_release(x_21, 1); - x_264 = x_21; -} else { - lean_dec_ref(x_21); - x_264 = lean_box(0); -} -x_265 = lean_ctor_get(x_157, 0); -lean_inc(x_265); -lean_dec(x_157); -x_266 = lean_array_push(x_6, x_265); -lean_inc(x_263); -if (lean_is_scalar(x_264)) { - x_267 = lean_alloc_ctor(0, 2, 0); -} else { - x_267 = x_264; -} -lean_ctor_set(x_267, 0, x_263); -lean_ctor_set(x_267, 1, x_18); -x_268 = lean_array_push(x_5, x_267); -x_269 = lean_nat_add(x_2, x_18); -x_270 = lean_box(0); -x_271 = l_List_tailD___rarg(x_3, x_270); -lean_dec(x_3); -x_272 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_272, 0, x_266); -lean_ctor_set(x_272, 1, x_268); -x_273 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_273, 0, x_271); -lean_ctor_set(x_273, 1, x_272); -x_274 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_274, 0, x_269); -lean_ctor_set(x_274, 1, x_273); -x_275 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_275, 0, x_7); -lean_ctor_set(x_275, 1, x_274); -x_276 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_276, 0, x_8); -lean_ctor_set(x_276, 1, x_275); -x_277 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_277, 0, x_263); -lean_ctor_set(x_277, 1, x_276); -x_278 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_278, 0, x_277); -lean_ctor_set(x_278, 1, x_262); -return x_278; -} -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_22, 0); +x_27 = lean_ctor_get(x_22, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_22); +x_28 = lean_array_push(x_13, x_26); +lean_ctor_set(x_3, 1, x_28); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_27); +return x_29; } } else { -uint8_t x_279; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_279 = !lean_is_exclusive(x_20); -if (x_279 == 0) +uint8_t x_30; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +x_30 = !lean_is_exclusive(x_22); +if (x_30 == 0) { -return x_20; +return x_22; } else { -lean_object* x_280; lean_object* x_281; lean_object* x_282; -x_280 = lean_ctor_get(x_20, 0); -x_281 = lean_ctor_get(x_20, 1); -lean_inc(x_281); -lean_inc(x_280); -lean_dec(x_20); -x_282 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_282, 0, x_280); -lean_ctor_set(x_282, 1, x_281); -return x_282; -} +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_22, 0); +x_32 = lean_ctor_get(x_22, 1); +lean_inc(x_32); +lean_inc(x_31); +lean_dec(x_22); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_11 = lean_ctor_get(x_3, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -x_16 = lean_ctor_get(x_3, 0); -lean_inc(x_16); -if (lean_is_exclusive(x_3)) { - lean_ctor_release(x_3, 0); - lean_ctor_release(x_3, 1); - x_17 = x_3; -} else { - lean_dec_ref(x_3); - x_17 = lean_box(0); -} -x_18 = lean_ctor_get(x_11, 0); -lean_inc(x_18); -if (lean_is_exclusive(x_11)) { - lean_ctor_release(x_11, 0); - lean_ctor_release(x_11, 1); - x_19 = x_11; -} else { - lean_dec_ref(x_11); - x_19 = lean_box(0); -} -x_20 = lean_ctor_get(x_12, 0); -lean_inc(x_20); -if (lean_is_exclusive(x_12)) { - lean_ctor_release(x_12, 0); - lean_ctor_release(x_12, 1); - x_21 = x_12; -} else { - lean_dec_ref(x_12); - x_21 = lean_box(0); -} -x_22 = lean_ctor_get(x_13, 0); -lean_inc(x_22); -if (lean_is_exclusive(x_13)) { - lean_ctor_release(x_13, 0); - lean_ctor_release(x_13, 1); - x_23 = x_13; -} else { - lean_dec_ref(x_13); - x_23 = lean_box(0); -} -x_24 = lean_ctor_get(x_14, 0); -lean_inc(x_24); -if (lean_is_exclusive(x_14)) { - lean_ctor_release(x_14, 0); - lean_ctor_release(x_14, 1); - x_25 = x_14; -} else { - lean_dec_ref(x_14); - x_25 = lean_box(0); -} -x_26 = lean_ctor_get(x_15, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_15, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_15)) { - lean_ctor_release(x_15, 0); - lean_ctor_release(x_15, 1); - x_28 = x_15; -} else { - lean_dec_ref(x_15); - x_28 = lean_box(0); +lean_object* x_34; lean_object* x_35; +x_34 = lean_array_fget(x_2, x_14); +lean_dec(x_14); +x_35 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(x_17, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_17); +if (lean_obj_tag(x_35) == 0) +{ +uint8_t x_36; +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); +x_38 = lean_array_push(x_13, x_37); +lean_ctor_set(x_3, 1, x_38); +lean_ctor_set(x_35, 0, x_3); +return x_35; } -lean_inc(x_22); -x_29 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_29, 0, x_22); -if (lean_obj_tag(x_2) == 0) +else { -lean_object* x_80; -x_80 = lean_box(0); -x_30 = x_80; -goto block_79; +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_35, 0); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_35); +x_41 = lean_array_push(x_13, x_39); +lean_ctor_set(x_3, 1, x_41); +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_3); +lean_ctor_set(x_42, 1, x_40); +return x_42; +} } else { -uint8_t x_81; -x_81 = !lean_is_exclusive(x_2); -if (x_81 == 0) +uint8_t x_43; +lean_free_object(x_3); +lean_dec(x_13); +lean_dec(x_12); +x_43 = !lean_is_exclusive(x_35); +if (x_43 == 0) { -lean_object* x_82; lean_object* x_83; -x_82 = lean_ctor_get(x_2, 0); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -lean_dec(x_82); -lean_ctor_set(x_2, 0, x_83); -x_30 = x_2; -goto block_79; +return x_35; } else { -lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_84 = lean_ctor_get(x_2, 0); -lean_inc(x_84); -lean_dec(x_2); -x_85 = lean_ctor_get(x_84, 0); -lean_inc(x_85); -lean_dec(x_84); -x_86 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_86, 0, x_85); -x_30 = x_86; -goto block_79; +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_35, 0); +x_45 = lean_ctor_get(x_35, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_35); +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; +} } } -block_79: -{ -uint8_t x_31; -x_31 = l___private_Init_Data_Option_Basic_0__Option_decEqOption____x40_Init_Data_Option_Basic___hyg_4____at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__1(x_29, x_30); -lean_dec(x_30); -lean_dec(x_29); -if (x_31 == 0) -{ -lean_object* x_32; uint8_t x_33; lean_object* x_34; -lean_dec(x_28); -lean_dec(x_25); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_17); -x_32 = lean_box(0); -x_33 = lean_unbox(x_16); -lean_dec(x_16); -x_34 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(x_1, x_22, x_24, x_33, x_27, x_26, x_20, x_18, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_22); -return x_34; } else { -uint8_t x_35; lean_object* x_36; lean_object* x_37; -lean_dec(x_20); -lean_dec(x_16); -x_35 = 0; -x_36 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_37 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_withoutParentProjections___rarg(x_35, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_37) == 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; uint8_t x_54; +x_47 = lean_ctor_get(x_3, 0); +x_48 = lean_ctor_get(x_3, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_3); +x_49 = lean_array_get_size(x_48); +x_50 = lean_nat_sub(x_1, x_49); +x_51 = lean_unsigned_to_nat(1u); +x_52 = lean_nat_sub(x_50, x_51); +lean_dec(x_50); +x_53 = lean_array_get_size(x_2); +x_54 = lean_nat_dec_lt(x_49, x_53); +lean_dec(x_53); +if (x_54 == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +lean_object* x_55; lean_object* x_56; lean_object* x_57; +lean_dec(x_49); +x_55 = l_Lean_PrettyPrinter_Delaborator_instInhabitedParamKind; +x_56 = l___private_Init_GetElem_0__outOfBounds___rarg(x_55); +x_57 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(x_52, x_56, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_52); +if (lean_obj_tag(x_57) == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_39 = lean_ctor_get(x_37, 0); -x_40 = lean_array_get_size(x_26); -x_41 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_41, 0, x_40); -x_42 = lean_unsigned_to_nat(1u); -x_43 = lean_nat_add(x_22, x_42); -lean_dec(x_22); -x_44 = lean_box(0); -x_45 = l_List_tailD___rarg(x_24, x_44); -lean_dec(x_24); -x_46 = lean_array_push(x_26, x_39); -x_47 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; -x_48 = lean_array_push(x_27, x_47); -if (lean_is_scalar(x_28)) { - x_49 = lean_alloc_ctor(0, 2, 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; +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_60 = x_57; } else { - x_49 = x_28; + lean_dec_ref(x_57); + x_60 = lean_box(0); } -lean_ctor_set(x_49, 0, x_46); -lean_ctor_set(x_49, 1, x_48); -if (lean_is_scalar(x_25)) { - x_50 = lean_alloc_ctor(0, 2, 0); +x_61 = lean_array_push(x_48, x_58); +x_62 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_62, 0, x_47); +lean_ctor_set(x_62, 1, x_61); +if (lean_is_scalar(x_60)) { + x_63 = lean_alloc_ctor(0, 2, 0); } else { - x_50 = x_25; + x_63 = x_60; } -lean_ctor_set(x_50, 0, x_45); -lean_ctor_set(x_50, 1, x_49); -if (lean_is_scalar(x_23)) { - x_51 = lean_alloc_ctor(0, 2, 0); -} else { - x_51 = x_23; +lean_ctor_set(x_63, 0, x_62); +lean_ctor_set(x_63, 1, x_59); +return x_63; } -lean_ctor_set(x_51, 0, x_43); -lean_ctor_set(x_51, 1, x_50); -if (lean_is_scalar(x_21)) { - x_52 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +lean_dec(x_48); +lean_dec(x_47); +x_64 = lean_ctor_get(x_57, 0); +lean_inc(x_64); +x_65 = lean_ctor_get(x_57, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_57)) { + lean_ctor_release(x_57, 0); + lean_ctor_release(x_57, 1); + x_66 = x_57; } else { - x_52 = x_21; + lean_dec_ref(x_57); + x_66 = lean_box(0); } -lean_ctor_set(x_52, 0, x_41); -lean_ctor_set(x_52, 1, x_51); -if (lean_is_scalar(x_19)) { - x_53 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_66)) { + x_67 = lean_alloc_ctor(1, 2, 0); } else { - x_53 = x_19; + x_67 = x_66; } -lean_ctor_set(x_53, 0, x_18); -lean_ctor_set(x_53, 1, x_52); -x_54 = lean_box(x_35); -if (lean_is_scalar(x_17)) { - x_55 = lean_alloc_ctor(0, 2, 0); -} else { - x_55 = x_17; +lean_ctor_set(x_67, 0, x_64); +lean_ctor_set(x_67, 1, x_65); +return x_67; } -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -lean_ctor_set(x_37, 0, x_55); -return x_37; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_56 = lean_ctor_get(x_37, 0); -x_57 = lean_ctor_get(x_37, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_37); -x_58 = lean_array_get_size(x_26); -x_59 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_59, 0, x_58); -x_60 = lean_unsigned_to_nat(1u); -x_61 = lean_nat_add(x_22, x_60); -lean_dec(x_22); -x_62 = lean_box(0); -x_63 = l_List_tailD___rarg(x_24, x_62); -lean_dec(x_24); -x_64 = lean_array_push(x_26, x_56); -x_65 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1; -x_66 = lean_array_push(x_27, x_65); -if (lean_is_scalar(x_28)) { - x_67 = lean_alloc_ctor(0, 2, 0); +lean_object* x_68; lean_object* x_69; +x_68 = lean_array_fget(x_2, x_49); +lean_dec(x_49); +x_69 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg(x_52, x_68, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_52); +if (lean_obj_tag(x_69) == 0) +{ +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_72 = x_69; } else { - x_67 = x_28; + lean_dec_ref(x_69); + x_72 = lean_box(0); } -lean_ctor_set(x_67, 0, x_64); -lean_ctor_set(x_67, 1, x_66); -if (lean_is_scalar(x_25)) { - x_68 = lean_alloc_ctor(0, 2, 0); +x_73 = lean_array_push(x_48, x_70); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_47); +lean_ctor_set(x_74, 1, x_73); +if (lean_is_scalar(x_72)) { + x_75 = lean_alloc_ctor(0, 2, 0); } else { - x_68 = x_25; + x_75 = x_72; } -lean_ctor_set(x_68, 0, x_63); -lean_ctor_set(x_68, 1, x_67); -if (lean_is_scalar(x_23)) { - x_69 = lean_alloc_ctor(0, 2, 0); -} else { - x_69 = x_23; +lean_ctor_set(x_75, 0, x_74); +lean_ctor_set(x_75, 1, x_71); +return x_75; } -lean_ctor_set(x_69, 0, x_61); -lean_ctor_set(x_69, 1, x_68); -if (lean_is_scalar(x_21)) { - x_70 = lean_alloc_ctor(0, 2, 0); +else +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +lean_dec(x_48); +lean_dec(x_47); +x_76 = lean_ctor_get(x_69, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_69, 1); +lean_inc(x_77); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_78 = x_69; } else { - x_70 = x_21; + lean_dec_ref(x_69); + x_78 = lean_box(0); } -lean_ctor_set(x_70, 0, x_59); -lean_ctor_set(x_70, 1, x_69); -if (lean_is_scalar(x_19)) { - x_71 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_78)) { + x_79 = lean_alloc_ctor(1, 2, 0); } else { - x_71 = x_19; + x_79 = x_78; } -lean_ctor_set(x_71, 0, x_18); -lean_ctor_set(x_71, 1, x_70); -x_72 = lean_box(x_35); -if (lean_is_scalar(x_17)) { - x_73 = lean_alloc_ctor(0, 2, 0); -} else { - x_73 = x_17; +lean_ctor_set(x_79, 0, x_76); +lean_ctor_set(x_79, 1, x_77); +return x_79; } -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_71); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_57); -return x_74; } } -else +} +} +LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3(lean_object* x_1) { +_start: { -uint8_t x_75; -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_19); -lean_dec(x_18); -lean_dec(x_17); -x_75 = !lean_is_exclusive(x_37); -if (x_75 == 0) +if (lean_obj_tag(x_1) == 0) { -return x_37; +uint8_t x_2; +x_2 = 0; +return x_2; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_37, 0); -x_77 = lean_ctor_get(x_37, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_37); -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; -} -} -} +uint8_t x_3; +x_3 = 1; +return x_3; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_12; uint8_t x_13; lean_object* x_14; -x_12 = l_Lean_Syntax_mkApp(x_1, x_2); -if (x_3 == 0) -{ -x_13 = x_3; -x_14 = x_11; -goto block_56; -} -else -{ -lean_object* x_57; lean_object* x_58; -x_57 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___closed__2; -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_58 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_57, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_58) == 0) -{ -lean_object* x_59; lean_object* x_60; uint8_t x_61; -x_59 = lean_ctor_get(x_58, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); -lean_dec(x_58); -x_61 = lean_unbox(x_59); -lean_dec(x_59); -x_13 = x_61; -x_14 = x_60; -goto block_56; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_11 = lean_array_get_size(x_1); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(x_1, x_12, x_11); +x_14 = l_Lean_Syntax_mkApp(x_2, 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; } -else -{ -uint8_t x_62; -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_62 = !lean_is_exclusive(x_58); -if (x_62 == 0) -{ -return x_58; } -else +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1() { +_start: { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_58, 0); -x_64 = lean_ctor_get(x_58, 1); -lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_58); -x_65 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_65, 0, x_63); -lean_ctor_set(x_65, 1, x_64); -return x_65; -} +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___boxed), 1, 0); +return x_1; } } -block_56: +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { +_start: { -if (x_13 == 0) +if (lean_obj_tag(x_6) == 0) { -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_object* x_14; lean_object* x_15; lean_dec(x_5); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_12); -lean_ctor_set(x_15, 1, x_14); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_14 = lean_box(0); +x_15 = lean_apply_8(x_1, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); return x_15; } else { -lean_object* x_16; lean_object* x_17; lean_object* x_22; -lean_inc(x_9); -lean_inc(x_12); -x_22 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance(x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_22) == 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; uint8_t x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +lean_dec(x_1); +x_16 = lean_ctor_get(x_6, 0); +lean_inc(x_16); +lean_dec(x_6); +x_17 = lean_unsigned_to_nat(0u); +lean_inc(x_3); +lean_inc(x_2); +x_18 = l_Array_toSubarray___rarg(x_2, x_17, x_3); +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_18, 2); +lean_inc(x_21); +lean_dec(x_18); +x_22 = lean_nat_dec_lt(x_20, x_21); +x_23 = lean_box(0); +x_24 = lean_array_set(x_2, x_3, x_23); +lean_dec(x_3); +if (x_22 == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -lean_dec(x_9); -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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_25, 0, x_23); -x_16 = x_25; -x_17 = x_24; -goto block_21; +uint8_t x_62; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +x_62 = 1; +x_25 = x_62; +goto block_61; } else { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_22); -if (x_26 == 0) -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_27 = lean_ctor_get(x_22, 0); -x_28 = lean_ctor_get(x_22, 1); -x_29 = l_Lean_Exception_isRuntime(x_27); -if (x_29 == 0) +lean_object* x_63; uint8_t x_64; +x_63 = lean_array_get_size(x_19); +x_64 = lean_nat_dec_le(x_21, x_63); +if (x_64 == 0) { -lean_dec(x_9); -if (lean_obj_tag(x_27) == 0) +uint8_t x_65; +lean_dec(x_21); +x_65 = lean_nat_dec_lt(x_20, x_63); +if (x_65 == 0) { -lean_dec(x_12); -return x_22; +uint8_t x_66; +lean_dec(x_63); +lean_dec(x_20); +lean_dec(x_19); +x_66 = 1; +x_25 = x_66; +goto block_61; } else { -lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_30 = lean_ctor_get(x_27, 0); -lean_inc(x_30); -x_31 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_32 = lean_nat_dec_eq(x_31, x_30); -lean_dec(x_30); -if (x_32 == 0) +size_t x_67; size_t x_68; uint8_t x_69; +x_67 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_68 = lean_usize_of_nat(x_63); +lean_dec(x_63); +x_69 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7(x_19, x_67, x_68); +lean_dec(x_19); +if (x_69 == 0) { -lean_dec(x_12); -return x_22; +uint8_t x_70; +x_70 = 1; +x_25 = x_70; +goto block_61; } else { -lean_object* x_33; -lean_free_object(x_22); -lean_dec(x_27); -x_33 = lean_box(0); -x_16 = x_33; -x_17 = x_28; -goto block_21; +uint8_t x_71; +x_71 = 0; +x_25 = x_71; +goto block_61; } } } else { -uint8_t x_34; -x_34 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -lean_dec(x_9); -if (x_34 == 0) +size_t x_72; size_t x_73; uint8_t x_74; +lean_dec(x_63); +x_72 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_73 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_74 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7(x_19, x_72, x_73); +lean_dec(x_19); +if (x_74 == 0) { -lean_dec(x_12); -return x_22; +uint8_t x_75; +x_75 = 1; +x_25 = x_75; +goto block_61; } else { -if (lean_obj_tag(x_27) == 0) -{ -lean_dec(x_12); -return x_22; +uint8_t x_76; +x_76 = 0; +x_25 = x_76; +goto block_61; } -else +} +} +block_61: { -lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_35 = lean_ctor_get(x_27, 0); -lean_inc(x_35); -x_36 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_37 = lean_nat_dec_eq(x_36, x_35); -lean_dec(x_35); -if (x_37 == 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_33; +x_26 = lean_ctor_get(x_11, 5); +lean_inc(x_26); +x_27 = 0; +x_28 = l_Lean_SourceInfo_fromRef(x_26, x_27); +x_29 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; +lean_inc(x_28); +x_30 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +x_31 = lean_mk_syntax_ident(x_4); +x_32 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; +x_33 = l_Lean_Syntax_node3(x_28, x_32, x_16, x_30, x_31); +if (x_25 == 0) { +lean_object* x_34; lean_object* x_35; +lean_dec(x_5); +x_34 = lean_box(0); +x_35 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_24, x_33, x_34, x_7, x_8, x_9, x_10, x_11, x_12, x_13); lean_dec(x_12); -return x_22; -} -else -{ -lean_object* x_38; -lean_free_object(x_22); -lean_dec(x_27); -x_38 = lean_box(0); -x_16 = x_38; -x_17 = x_28; -goto block_21; -} -} -} -} +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_24); +return x_35; } else { -lean_object* x_39; lean_object* x_40; uint8_t x_41; -x_39 = lean_ctor_get(x_22, 0); -x_40 = lean_ctor_get(x_22, 1); -lean_inc(x_40); -lean_inc(x_39); -lean_dec(x_22); -x_41 = l_Lean_Exception_isRuntime(x_39); -if (x_41 == 0) +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_37 = l_Array_findIdx_x3f_loop___rarg(x_24, x_36, x_17); +x_38 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_annotateTermInfo___boxed), 8, 1); +lean_closure_set(x_38, 0, x_33); +if (lean_obj_tag(x_37) == 0) { -lean_dec(x_9); -if (lean_obj_tag(x_39) == 0) +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_array_get_size(x_24); +x_40 = lean_nat_sub(x_5, x_39); +lean_dec(x_39); +lean_dec(x_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_41 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_40, x_38, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_42; +lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = lean_ctor_get(x_41, 0); +lean_inc(x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = lean_box(0); +x_45 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_24, x_42, x_44, x_7, x_8, x_9, x_10, x_11, x_12, x_43); lean_dec(x_12); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_39); -lean_ctor_set(x_42, 1, x_40); -return x_42; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_24); +return x_45; } else { -lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_43 = lean_ctor_get(x_39, 0); -lean_inc(x_43); -x_44 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_45 = lean_nat_dec_eq(x_44, x_43); -lean_dec(x_43); -if (x_45 == 0) -{ -lean_object* x_46; +uint8_t x_46; +lean_dec(x_24); lean_dec(x_12); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_39); -lean_ctor_set(x_46, 1, x_40); -return x_46; -} -else +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_46 = !lean_is_exclusive(x_41); +if (x_46 == 0) { -lean_object* x_47; -lean_dec(x_39); -x_47 = lean_box(0); -x_16 = x_47; -x_17 = x_40; -goto block_21; -} -} +return x_41; } else { -uint8_t x_48; -x_48 = lean_ctor_get_uint8(x_9, sizeof(void*)*11); -lean_dec(x_9); -if (x_48 == 0) -{ -lean_object* x_49; -lean_dec(x_12); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_41, 0); +x_48 = lean_ctor_get(x_41, 1); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_41); x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_39); -lean_ctor_set(x_49, 1, x_40); +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); return x_49; } +} +} else { -if (lean_obj_tag(x_39) == 0) +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_37, 0); +lean_inc(x_50); +lean_dec(x_37); +x_51 = lean_nat_sub(x_5, x_50); +lean_dec(x_50); +lean_dec(x_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_52 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_51, x_38, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_52) == 0) { -lean_object* x_50; +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = lean_box(0); +x_56 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_24, x_53, x_55, x_7, x_8, x_9, x_10, x_11, x_12, x_54); lean_dec(x_12); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_39); -lean_ctor_set(x_50, 1, x_40); -return x_50; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_24); +return x_56; } else { -lean_object* x_51; lean_object* x_52; uint8_t x_53; -x_51 = lean_ctor_get(x_39, 0); -lean_inc(x_51); -x_52 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_53 = lean_nat_dec_eq(x_52, x_51); -lean_dec(x_51); -if (x_53 == 0) -{ -lean_object* x_54; +uint8_t x_57; +lean_dec(x_24); lean_dec(x_12); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_39); -lean_ctor_set(x_54, 1, x_40); -return x_54; +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +x_57 = !lean_is_exclusive(x_52); +if (x_57 == 0) +{ +return x_52; } else { -lean_object* x_55; -lean_dec(x_39); -x_55 = lean_box(0); -x_16 = x_55; -x_17 = x_40; -goto block_21; +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_52, 0); +x_59 = lean_ctor_get(x_52, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_52); +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; } } } } } } -block_21: +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -if (lean_obj_tag(x_16) == 0) +lean_object* x_9; +x_9 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections(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_18; -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_12); -lean_ctor_set(x_18, 1, x_17); -return x_18; +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; } else { -lean_object* x_19; lean_object* x_20; -lean_dec(x_12); -x_19 = lean_ctor_get(x_16, 0); -lean_inc(x_19); -lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -return x_20; -} +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_15, 0, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; } } +else +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +return x_9; } +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_getPPNotation___boxed), 1, 0); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { _start: { -lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; -lean_dec(x_5); -x_13 = lean_array_get_size(x_1); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_nat_dec_lt(x_14, x_13); -if (x_15 == 0) +lean_dec(x_6); +if (lean_obj_tag(x_2) == 0) { -uint8_t x_62; -lean_dec(x_13); -x_62 = 0; -x_16 = x_62; -x_17 = x_12; -goto block_61; +lean_object* x_14; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_1); +lean_ctor_set(x_14, 1, x_13); +return x_14; } else { -size_t x_63; size_t x_64; uint8_t x_65; -x_63 = 0; -x_64 = lean_usize_of_nat(x_13); -lean_dec(x_13); -x_65 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(x_1, x_63, x_64); -if (x_65 == 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_2, 0); +lean_inc(x_15); +lean_dec(x_2); +lean_inc(x_1); +x_16 = lean_alloc_closure((void*)(l_Lean_unresolveNameGlobal_unresolveNameCore___at_Lean_PrettyPrinter_Delaborator_delabConst___spec__2___lambda__1___boxed), 9, 1); +lean_closure_set(x_16, 0, x_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); +lean_dec(x_15); +x_19 = lean_nat_dec_lt(x_17, x_3); +lean_dec(x_3); +if (x_19 == 0) { -x_16 = x_65; -x_17 = x_12; -goto block_61; +lean_object* x_20; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_1); +lean_ctor_set(x_20, 1, x_13); +return x_20; } else { -lean_object* x_66; lean_object* x_67; -x_66 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +lean_object* x_21; +lean_dec(x_1); +x_21 = lean_array_fget(x_4, x_17); +if (lean_obj_tag(x_21) == 1) +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6), 8, 1); +lean_closure_set(x_23, 0, x_22); +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_67 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_66, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_67) == 0) +x_24 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__8(x_17, x_23, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_68; lean_object* x_69; uint8_t 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); -x_70 = lean_unbox(x_68); -lean_dec(x_68); -x_16 = x_70; -x_17 = x_69; -goto block_61; +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_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(x_16, x_4, x_17, x_18, x_5, x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +return x_27; } else { -uint8_t x_71; +uint8_t x_28; +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_16); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) +{ +return x_24; +} +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; +} +} +} +else +{ +lean_object* x_32; lean_object* x_33; +lean_dec(x_21); +x_32 = lean_box(0); +x_33 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(x_16, x_4, x_17, x_18, x_5, x_32, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_33; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__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, 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; uint8_t x_18; lean_object* x_19; lean_object* x_64; lean_dec(x_6); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_71 = !lean_is_exclusive(x_67); -if (x_71 == 0) +x_14 = lean_array_get_size(x_1); +x_15 = lean_unsigned_to_nat(0u); +lean_inc(x_14); +x_16 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(x_1, x_15, x_14); +x_17 = l_Lean_Syntax_mkApp(x_2, x_16); +if (x_5 == 0) { -return x_67; +uint8_t x_75; +x_75 = 0; +x_18 = x_75; +x_19 = x_13; +goto block_63; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_67, 0); -x_73 = lean_ctor_get(x_67, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_67); -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; +uint8_t x_76; +x_76 = lean_nat_dec_lt(x_15, x_14); +if (x_76 == 0) +{ +lean_object* x_77; +x_77 = lean_box(0); +x_64 = x_77; +goto block_74; +} +else +{ +size_t x_78; size_t x_79; uint8_t x_80; +x_78 = 0; +x_79 = lean_usize_of_nat(x_14); +x_80 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9(x_1, x_78, x_79); +if (x_80 == 0) +{ +lean_object* x_81; +x_81 = lean_box(0); +x_64 = x_81; +goto block_74; } +else +{ +uint8_t x_82; +x_82 = 0; +x_18 = x_82; +x_19 = x_13; +goto block_63; } } } -block_61: +block_63: { -if (x_16 == 0) +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; -lean_dec(x_1); -x_18 = lean_box(0); -x_19 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_2, x_3, x_4, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_17); -return x_19; +lean_object* x_20; lean_object* x_21; +x_20 = lean_box(0); +x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7(x_17, x_3, x_14, x_1, x_4, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +return x_21; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_27; +lean_object* x_22; lean_object* x_23; lean_object* x_29; +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_3); -lean_inc(x_2); -x_27 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(x_2, x_3, x_1, x_6, x_7, x_8, x_9, x_10, x_11, x_17); -if (lean_obj_tag(x_27) == 0) +lean_inc(x_17); +x_29 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance(x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_19); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_28); -x_20 = x_30; -x_21 = x_29; -goto block_26; +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 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_32, 0, x_30); +x_22 = x_32; +x_23 = x_31; +goto block_28; } else { -uint8_t x_31; -x_31 = !lean_is_exclusive(x_27); -if (x_31 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_29); +if (x_33 == 0) { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_32 = lean_ctor_get(x_27, 0); -x_33 = lean_ctor_get(x_27, 1); -x_34 = l_Lean_Exception_isRuntime(x_32); -if (x_34 == 0) +lean_object* x_34; lean_object* x_35; uint8_t x_36; +x_34 = lean_ctor_get(x_29, 0); +x_35 = lean_ctor_get(x_29, 1); +x_36 = l_Lean_Exception_isRuntime(x_34); +if (x_36 == 0) { -if (lean_obj_tag(x_32) == 0) +if (lean_obj_tag(x_34) == 0) { +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_27; +lean_dec(x_1); +return x_29; } else { -lean_object* x_35; lean_object* x_36; uint8_t x_37; -x_35 = lean_ctor_get(x_32, 0); -lean_inc(x_35); -x_36 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_37 = lean_nat_dec_eq(x_36, x_35); -lean_dec(x_35); -if (x_37 == 0) +lean_object* x_37; lean_object* x_38; uint8_t x_39; +x_37 = lean_ctor_get(x_34, 0); +lean_inc(x_37); +x_38 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_39 = lean_nat_dec_eq(x_38, x_37); +lean_dec(x_37); +if (x_39 == 0) { +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_27; +lean_dec(x_1); +return x_29; } else { -lean_object* x_38; -lean_free_object(x_27); -lean_dec(x_32); -x_38 = lean_box(0); -x_20 = x_38; -x_21 = x_33; -goto block_26; +lean_object* x_40; +lean_free_object(x_29); +lean_dec(x_34); +x_40 = lean_box(0); +x_22 = x_40; +x_23 = x_35; +goto block_28; } } } else { -uint8_t x_39; -x_39 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -if (x_39 == 0) +uint8_t x_41; +x_41 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); +if (x_41 == 0) { +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_27; +lean_dec(x_1); +return x_29; } else { -if (lean_obj_tag(x_32) == 0) +if (lean_obj_tag(x_34) == 0) { +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_27; +lean_dec(x_1); +return x_29; } else { -lean_object* x_40; lean_object* x_41; uint8_t x_42; -x_40 = lean_ctor_get(x_32, 0); -lean_inc(x_40); -x_41 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_42 = lean_nat_dec_eq(x_41, x_40); -lean_dec(x_40); -if (x_42 == 0) +lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_42 = lean_ctor_get(x_34, 0); +lean_inc(x_42); +x_43 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_44 = lean_nat_dec_eq(x_43, x_42); +lean_dec(x_42); +if (x_44 == 0) { +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -return x_27; +lean_dec(x_1); +return x_29; } else { -lean_object* x_43; -lean_free_object(x_27); -lean_dec(x_32); -x_43 = lean_box(0); -x_20 = x_43; -x_21 = x_33; -goto block_26; +lean_object* x_45; +lean_free_object(x_29); +lean_dec(x_34); +x_45 = lean_box(0); +x_22 = x_45; +x_23 = x_35; +goto block_28; } } } @@ -17990,657 +17548,534 @@ goto block_26; } else { -lean_object* x_44; lean_object* x_45; uint8_t 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 = l_Lean_Exception_isRuntime(x_44); -if (x_46 == 0) +lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_46 = lean_ctor_get(x_29, 0); +x_47 = lean_ctor_get(x_29, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_29); +x_48 = l_Lean_Exception_isRuntime(x_46); +if (x_48 == 0) { -if (lean_obj_tag(x_44) == 0) +if (lean_obj_tag(x_46) == 0) { -lean_object* x_47; +lean_object* x_49; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_47 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_47, 0, x_44); -lean_ctor_set(x_47, 1, x_45); -return x_47; +lean_dec(x_1); +x_49 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_49, 0, x_46); +lean_ctor_set(x_49, 1, x_47); +return x_49; } else { -lean_object* x_48; lean_object* x_49; uint8_t x_50; -x_48 = lean_ctor_get(x_44, 0); -lean_inc(x_48); -x_49 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_50 = lean_nat_dec_eq(x_49, x_48); -lean_dec(x_48); -if (x_50 == 0) +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = lean_ctor_get(x_46, 0); +lean_inc(x_50); +x_51 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_52 = lean_nat_dec_eq(x_51, x_50); +lean_dec(x_50); +if (x_52 == 0) { -lean_object* x_51; +lean_object* x_53; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_51 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_51, 0, x_44); -lean_ctor_set(x_51, 1, x_45); -return x_51; +lean_dec(x_1); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_46); +lean_ctor_set(x_53, 1, x_47); +return x_53; } else { -lean_object* x_52; -lean_dec(x_44); -x_52 = lean_box(0); -x_20 = x_52; -x_21 = x_45; -goto block_26; +lean_object* x_54; +lean_dec(x_46); +x_54 = lean_box(0); +x_22 = x_54; +x_23 = x_47; +goto block_28; } } } else { -uint8_t x_53; -x_53 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); -if (x_53 == 0) +uint8_t x_55; +x_55 = lean_ctor_get_uint8(x_11, sizeof(void*)*11); +if (x_55 == 0) { -lean_object* x_54; +lean_object* x_56; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_54 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_54, 0, x_44); -lean_ctor_set(x_54, 1, x_45); -return x_54; +lean_dec(x_1); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_46); +lean_ctor_set(x_56, 1, x_47); +return x_56; } else { -if (lean_obj_tag(x_44) == 0) +if (lean_obj_tag(x_46) == 0) { -lean_object* x_55; +lean_object* x_57; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_44); -lean_ctor_set(x_55, 1, x_45); -return x_55; +lean_dec(x_1); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_46); +lean_ctor_set(x_57, 1, x_47); +return x_57; } else { -lean_object* x_56; lean_object* x_57; uint8_t x_58; -x_56 = lean_ctor_get(x_44, 0); -lean_inc(x_56); -x_57 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; -x_58 = lean_nat_dec_eq(x_57, x_56); -lean_dec(x_56); -if (x_58 == 0) +lean_object* x_58; lean_object* x_59; uint8_t x_60; +x_58 = lean_ctor_get(x_46, 0); +lean_inc(x_58); +x_59 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_60 = lean_nat_dec_eq(x_59, x_58); +lean_dec(x_58); +if (x_60 == 0) { -lean_object* x_59; +lean_object* x_61; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_44); -lean_ctor_set(x_59, 1, x_45); -return x_59; +lean_dec(x_1); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_46); +lean_ctor_set(x_61, 1, x_47); +return x_61; } else { -lean_object* x_60; -lean_dec(x_44); -x_60 = lean_box(0); -x_20 = x_60; -x_21 = x_45; -goto block_26; +lean_object* x_62; +lean_dec(x_46); +x_62 = lean_box(0); +x_22 = x_62; +x_23 = x_47; +goto block_28; } } } } } } -block_26: +block_28: { -if (lean_obj_tag(x_20) == 0) +if (lean_obj_tag(x_22) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_box(0); -x_23 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_2, x_3, x_4, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -return x_23; +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7(x_17, x_3, x_14, x_1, x_4, x_24, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +return x_25; } else { -lean_object* x_24; lean_object* x_25; +lean_object* x_26; lean_object* x_27; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -lean_dec(x_6); +lean_dec(x_4); lean_dec(x_3); -lean_dec(x_2); -x_24 = lean_ctor_get(x_20, 0); -lean_inc(x_24); -lean_dec(x_20); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_21); -return x_25; +lean_dec(x_1); +x_26 = lean_ctor_get(x_22, 0); +lean_inc(x_26); +lean_dec(x_22); +x_27 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_23); +return x_27; } } } } +block_74: +{ +lean_object* x_65; lean_object* x_66; +lean_dec(x_64); +x_65 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___closed__2; +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_66 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_65, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_66) == 0) +{ +lean_object* x_67; lean_object* x_68; uint8_t x_69; +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_unbox(x_67); +lean_dec(x_67); +x_18 = x_69; +x_19 = x_68; +goto block_63; } +else +{ +uint8_t x_70; +lean_dec(x_17); +lean_dec(x_14); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +x_70 = !lean_is_exclusive(x_66); +if (x_70 == 0) +{ +return x_66; } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); -return x_1; +lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_71 = lean_ctor_get(x_66, 0); +x_72 = lean_ctor_get(x_66, 1); +lean_inc(x_72); +lean_inc(x_71); +lean_dec(x_66); +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; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Option.get!", 11); -return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3() { +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("value is none", 13); +x_1 = lean_alloc_closure((void*)(l_Lean_getPPNotation___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___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_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6(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_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_13 = lean_box(x_3); -lean_inc(x_2); -x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1___boxed), 11, 3); -lean_closure_set(x_14, 0, x_1); -lean_closure_set(x_14, 1, x_2); -lean_closure_set(x_14, 2, x_13); -x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); -lean_closure_set(x_15, 0, x_4); -lean_closure_set(x_15, 1, x_14); -lean_inc(x_5); -lean_inc(x_2); -x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___boxed), 10, 2); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_5); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_inc(x_1); +x_13 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1___boxed), 9, 1); +lean_closure_set(x_13, 0, x_1); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); +lean_closure_set(x_14, 0, x_2); +lean_closure_set(x_14, 1, x_13); +lean_inc(x_1); +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___boxed), 10, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_3); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_2); -x_17 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__1(x_2, x_15, x_16, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_17) == 0) +lean_inc(x_1); +x_16 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__1(x_1, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -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; uint8_t x_22; lean_object* x_23; +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_18, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_19, 1); +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + lean_ctor_release(x_16, 1); + x_19 = x_16; +} else { + lean_dec_ref(x_16); + x_19 = lean_box(0); +} +x_20 = lean_ctor_get(x_17, 0); lean_inc(x_20); -x_21 = lean_ctor_get(x_20, 1); +x_21 = lean_ctor_get(x_17, 1); lean_inc(x_21); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_ctor_get(x_20, 0); -lean_inc(x_24); -lean_dec(x_20); -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; uint8_t x_31; lean_object* x_32; -lean_dec(x_5); -lean_dec(x_2); -x_25 = lean_ctor_get(x_17, 1); -lean_inc(x_25); -lean_dec(x_17); -x_26 = lean_ctor_get(x_18, 0); -lean_inc(x_26); -lean_dec(x_18); -x_27 = lean_ctor_get(x_19, 0); -lean_inc(x_27); -lean_dec(x_19); -x_28 = lean_ctor_get(x_23, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_23, 1); -lean_inc(x_29); -lean_dec(x_23); -x_30 = lean_box(0); -x_31 = lean_unbox(x_26); -lean_dec(x_26); -x_32 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(x_29, x_27, x_28, x_31, x_30, x_6, x_7, x_8, x_9, x_10, x_11, x_25); -return x_32; -} -else -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; -lean_dec(x_19); -lean_dec(x_18); -x_33 = lean_ctor_get(x_17, 1); -lean_inc(x_33); lean_dec(x_17); -x_34 = lean_ctor_get(x_23, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_23, 1); -lean_inc(x_35); -lean_dec(x_23); -x_36 = lean_ctor_get(x_24, 0); -lean_inc(x_36); -lean_dec(x_24); -x_37 = lean_array_get_size(x_34); -x_38 = lean_nat_dec_lt(x_36, x_37); -lean_dec(x_37); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_192; lean_object* x_193; -x_192 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4; -x_193 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(x_192); -x_39 = x_193; -goto block_191; -} -else -{ -lean_object* x_194; -x_194 = lean_ctor_get(x_5, 0); -lean_inc(x_194); -lean_dec(x_5); -x_39 = x_194; -goto block_191; -} -block_191: -{ -if (x_38 == 0) -{ -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_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_40 = lean_ctor_get(x_39, 1); -lean_inc(x_40); -lean_dec(x_39); -x_41 = lean_box(0); -x_42 = l___private_Init_GetElem_0__outOfBounds___rarg(x_41); -x_43 = lean_ctor_get(x_10, 5); -lean_inc(x_43); -x_44 = 0; -x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); -x_46 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; -lean_inc(x_45); -x_47 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_47, 0, x_45); -lean_ctor_set(x_47, 1, x_46); -x_48 = lean_mk_syntax_ident(x_40); -x_49 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2; -x_50 = l_Lean_Syntax_node3(x_45, x_49, x_42, x_47, x_48); -x_51 = lean_unsigned_to_nat(0u); -x_52 = lean_nat_dec_eq(x_36, x_51); -if (x_52 == 0) +if (x_4 == 0) { -lean_object* x_53; lean_object* x_54; -lean_dec(x_35); -lean_dec(x_2); -x_53 = lean_box(0); -x_54 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_50, x_53, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_54; +x_22 = x_4; +x_23 = x_18; +goto block_67; } else { -lean_object* x_55; uint8_t x_56; lean_object* x_57; uint8_t x_58; lean_object* x_59; -x_55 = lean_array_get_size(x_35); -x_56 = lean_nat_dec_lt(x_51, x_55); -x_57 = lean_unsigned_to_nat(1u); -x_58 = lean_nat_dec_lt(x_57, x_55); -lean_dec(x_55); -x_59 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_annotateTermInfo___boxed), 8, 1); -lean_closure_set(x_59, 0, x_50); -if (x_56 == 0) -{ -lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_60 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_61 = l___private_Init_GetElem_0__outOfBounds___rarg(x_60); -x_62 = lean_ctor_get(x_61, 1); -lean_inc(x_62); -lean_dec(x_61); -x_63 = lean_nat_sub(x_2, x_62); -lean_dec(x_62); -lean_dec(x_2); -if (x_58 == 0) -{ -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -lean_dec(x_35); -x_64 = l___private_Init_GetElem_0__outOfBounds___rarg(x_60); -x_65 = lean_ctor_get(x_64, 1); -lean_inc(x_65); -lean_dec(x_64); -x_66 = lean_nat_sub(x_63, x_65); -lean_dec(x_65); -lean_dec(x_63); +lean_object* x_68; lean_object* x_69; +x_68 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__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_67 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_66, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_67) == 0) +x_69 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_68, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_70 = lean_box(0); -x_71 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_68, x_70, x_6, x_7, x_8, x_9, x_10, x_11, x_69); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_71; +lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = lean_unbox(x_70); +lean_dec(x_70); +x_22 = x_72; +x_23 = x_71; +goto block_67; } else { -uint8_t x_72; -lean_dec(x_36); -lean_dec(x_34); +uint8_t x_73; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_72 = !lean_is_exclusive(x_67); -if (x_72 == 0) +lean_dec(x_5); +lean_dec(x_1); +x_73 = !lean_is_exclusive(x_69); +if (x_73 == 0) { -return x_67; +return x_69; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; -x_73 = lean_ctor_get(x_67, 0); -x_74 = lean_ctor_get(x_67, 1); +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_69, 0); +x_75 = lean_ctor_get(x_69, 1); +lean_inc(x_75); lean_inc(x_74); -lean_inc(x_73); -lean_dec(x_67); -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_dec(x_69); +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; +} } } +block_67: +{ +if (x_22 == 0) +{ +lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_24 = lean_box(0); +x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__8(x_21, x_20, x_5, x_1, x_4, x_24, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +return x_25; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_76 = lean_array_fget(x_35, x_57); -lean_dec(x_35); -x_77 = lean_ctor_get(x_76, 1); -lean_inc(x_77); -lean_dec(x_76); -x_78 = lean_nat_sub(x_63, x_77); -lean_dec(x_77); -lean_dec(x_63); +lean_object* x_26; lean_object* x_27; lean_object* x_33; 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_79 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_78, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_79) == 0) +lean_inc(x_21); +lean_inc(x_20); +x_33 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders(x_20, x_21, x_6, x_7, x_8, x_9, x_10, x_11, x_23); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -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_box(0); -x_83 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_80, x_82, x_6, x_7, x_8, x_9, x_10, x_11, x_81); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_36, 0, x_34); +x_26 = x_36; +x_27 = x_35; +goto block_32; +} +else +{ +uint8_t x_37; +x_37 = !lean_is_exclusive(x_33); +if (x_37 == 0) +{ +lean_object* x_38; lean_object* x_39; uint8_t x_40; +x_38 = lean_ctor_get(x_33, 0); +x_39 = lean_ctor_get(x_33, 1); +x_40 = l_Lean_Exception_isRuntime(x_38); +if (x_40 == 0) +{ +if (lean_obj_tag(x_38) == 0) +{ +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -return x_83; +lean_dec(x_5); +lean_dec(x_1); +return x_33; } else { -uint8_t x_84; -lean_dec(x_36); -lean_dec(x_34); +lean_object* x_41; lean_object* x_42; uint8_t x_43; +x_41 = lean_ctor_get(x_38, 0); +lean_inc(x_41); +x_42 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_43 = lean_nat_dec_eq(x_42, x_41); +lean_dec(x_41); +if (x_43 == 0) +{ +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_84 = !lean_is_exclusive(x_79); -if (x_84 == 0) -{ -return x_79; +lean_dec(x_5); +lean_dec(x_1); +return x_33; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_85 = lean_ctor_get(x_79, 0); -x_86 = lean_ctor_get(x_79, 1); -lean_inc(x_86); -lean_inc(x_85); -lean_dec(x_79); -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_44; +lean_free_object(x_33); +lean_dec(x_38); +x_44 = lean_box(0); +x_26 = x_44; +x_27 = x_39; +goto block_32; } } } else { -lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_88 = lean_array_fget(x_35, x_51); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -x_90 = lean_nat_sub(x_2, x_89); -lean_dec(x_89); -lean_dec(x_2); -if (x_58 == 0) -{ -lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; -lean_dec(x_35); -x_91 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_92 = l___private_Init_GetElem_0__outOfBounds___rarg(x_91); -x_93 = lean_ctor_get(x_92, 1); -lean_inc(x_93); -lean_dec(x_92); -x_94 = lean_nat_sub(x_90, x_93); -lean_dec(x_93); -lean_dec(x_90); -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_95 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_94, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_95) == 0) +uint8_t x_45; +x_45 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); +if (x_45 == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -lean_dec(x_95); -x_98 = lean_box(0); -x_99 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_96, x_98, x_6, x_7, x_8, x_9, x_10, x_11, x_97); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -return x_99; +lean_dec(x_5); +lean_dec(x_1); +return x_33; } else { -uint8_t x_100; -lean_dec(x_36); -lean_dec(x_34); +if (lean_obj_tag(x_38) == 0) +{ +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_100 = !lean_is_exclusive(x_95); -if (x_100 == 0) -{ -return x_95; -} -else -{ -lean_object* x_101; lean_object* x_102; lean_object* x_103; -x_101 = lean_ctor_get(x_95, 0); -x_102 = lean_ctor_get(x_95, 1); -lean_inc(x_102); -lean_inc(x_101); -lean_dec(x_95); -x_103 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_103, 0, x_101); -lean_ctor_set(x_103, 1, x_102); -return x_103; -} -} +lean_dec(x_5); +lean_dec(x_1); +return x_33; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; -x_104 = lean_array_fget(x_35, x_57); -lean_dec(x_35); -x_105 = lean_ctor_get(x_104, 1); -lean_inc(x_105); -lean_dec(x_104); -x_106 = lean_nat_sub(x_90, x_105); -lean_dec(x_105); -lean_dec(x_90); -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_107 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_106, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_107) == 0) -{ -lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -x_109 = lean_ctor_get(x_107, 1); -lean_inc(x_109); -lean_dec(x_107); -x_110 = lean_box(0); -x_111 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_108, x_110, x_6, x_7, x_8, x_9, x_10, x_11, x_109); -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_111; -} -else +lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_46 = lean_ctor_get(x_38, 0); +lean_inc(x_46); +x_47 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_48 = lean_nat_dec_eq(x_47, x_46); +lean_dec(x_46); +if (x_48 == 0) { -uint8_t x_112; -lean_dec(x_36); -lean_dec(x_34); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_112 = !lean_is_exclusive(x_107); -if (x_112 == 0) -{ -return x_107; +lean_dec(x_5); +lean_dec(x_1); +return x_33; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_107, 0); -x_114 = lean_ctor_get(x_107, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_107); -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_49; +lean_free_object(x_33); +lean_dec(x_38); +x_49 = lean_box(0); +x_26 = x_49; +x_27 = x_39; +goto block_32; } } } @@ -18648,342 +18083,191 @@ return x_115; } else { -lean_object* x_116; lean_object* x_117; lean_object* x_118; uint8_t 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; uint8_t x_127; -x_116 = lean_ctor_get(x_39, 1); -lean_inc(x_116); -lean_dec(x_39); -x_117 = lean_array_fget(x_34, x_36); -x_118 = lean_ctor_get(x_10, 5); -lean_inc(x_118); -x_119 = 0; -x_120 = l_Lean_SourceInfo_fromRef(x_118, x_119); -x_121 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; -lean_inc(x_120); -x_122 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = lean_mk_syntax_ident(x_116); -x_124 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2; -x_125 = l_Lean_Syntax_node3(x_120, x_124, x_117, x_122, x_123); -x_126 = lean_unsigned_to_nat(0u); -x_127 = lean_nat_dec_eq(x_36, x_126); -if (x_127 == 0) -{ -lean_object* x_128; lean_object* x_129; -lean_dec(x_35); -lean_dec(x_2); -x_128 = lean_box(0); -x_129 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_125, x_128, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_129; -} -else -{ -lean_object* x_130; uint8_t x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; -x_130 = lean_array_get_size(x_35); -x_131 = lean_nat_dec_lt(x_126, x_130); -x_132 = lean_unsigned_to_nat(1u); -x_133 = lean_nat_dec_lt(x_132, x_130); -lean_dec(x_130); -x_134 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_annotateTermInfo___boxed), 8, 1); -lean_closure_set(x_134, 0, x_125); -if (x_131 == 0) -{ -lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_135 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_136 = l___private_Init_GetElem_0__outOfBounds___rarg(x_135); -x_137 = lean_ctor_get(x_136, 1); -lean_inc(x_137); -lean_dec(x_136); -x_138 = lean_nat_sub(x_2, x_137); -lean_dec(x_137); -lean_dec(x_2); -if (x_133 == 0) +lean_object* x_50; lean_object* x_51; uint8_t x_52; +x_50 = lean_ctor_get(x_33, 0); +x_51 = lean_ctor_get(x_33, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_33); +x_52 = l_Lean_Exception_isRuntime(x_50); +if (x_52 == 0) { -lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_dec(x_35); -x_139 = l___private_Init_GetElem_0__outOfBounds___rarg(x_135); -x_140 = lean_ctor_get(x_139, 1); -lean_inc(x_140); -lean_dec(x_139); -x_141 = lean_nat_sub(x_138, x_140); -lean_dec(x_140); -lean_dec(x_138); -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_142 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_141, x_134, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_142) == 0) +if (lean_obj_tag(x_50) == 0) { -lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_ctor_get(x_142, 1); -lean_inc(x_144); -lean_dec(x_142); -x_145 = lean_box(0); -x_146 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_143, x_145, x_6, x_7, x_8, x_9, x_10, x_11, x_144); +lean_object* x_53; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -return x_146; +lean_dec(x_5); +lean_dec(x_1); +x_53 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_53, 0, x_50); +lean_ctor_set(x_53, 1, x_51); +return x_53; } else { -uint8_t x_147; -lean_dec(x_36); -lean_dec(x_34); +lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_54 = lean_ctor_get(x_50, 0); +lean_inc(x_54); +x_55 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_56 = lean_nat_dec_eq(x_55, x_54); +lean_dec(x_54); +if (x_56 == 0) +{ +lean_object* x_57; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_147 = !lean_is_exclusive(x_142); -if (x_147 == 0) -{ -return x_142; +lean_dec(x_5); +lean_dec(x_1); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_50); +lean_ctor_set(x_57, 1, x_51); +return x_57; } else { -lean_object* x_148; lean_object* x_149; lean_object* x_150; -x_148 = lean_ctor_get(x_142, 0); -x_149 = lean_ctor_get(x_142, 1); -lean_inc(x_149); -lean_inc(x_148); -lean_dec(x_142); -x_150 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_150, 0, x_148); -lean_ctor_set(x_150, 1, x_149); -return x_150; +lean_object* x_58; +lean_dec(x_50); +x_58 = lean_box(0); +x_26 = x_58; +x_27 = x_51; +goto block_32; } } } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -x_151 = lean_array_fget(x_35, x_132); -lean_dec(x_35); -x_152 = lean_ctor_get(x_151, 1); -lean_inc(x_152); -lean_dec(x_151); -x_153 = lean_nat_sub(x_138, x_152); -lean_dec(x_152); -lean_dec(x_138); -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_154 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_153, x_134, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_154) == 0) +uint8_t x_59; +x_59 = lean_ctor_get_uint8(x_10, sizeof(void*)*11); +if (x_59 == 0) { -lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); -lean_inc(x_156); -lean_dec(x_154); -x_157 = lean_box(0); -x_158 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_155, x_157, x_6, x_7, x_8, x_9, x_10, x_11, x_156); +lean_object* x_60; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -return x_158; +lean_dec(x_5); +lean_dec(x_1); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_50); +lean_ctor_set(x_60, 1, x_51); +return x_60; } else { -uint8_t x_159; -lean_dec(x_36); -lean_dec(x_34); +if (lean_obj_tag(x_50) == 0) +{ +lean_object* x_61; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -x_159 = !lean_is_exclusive(x_154); -if (x_159 == 0) -{ -return x_154; -} -else -{ -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_154, 0); -x_161 = lean_ctor_get(x_154, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_154); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); -return x_162; -} -} -} +lean_dec(x_5); +lean_dec(x_1); +x_61 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_61, 0, x_50); +lean_ctor_set(x_61, 1, x_51); +return x_61; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = lean_array_fget(x_35, x_126); -x_164 = lean_ctor_get(x_163, 1); -lean_inc(x_164); -lean_dec(x_163); -x_165 = lean_nat_sub(x_2, x_164); -lean_dec(x_164); -lean_dec(x_2); -if (x_133 == 0) -{ -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_35); -x_166 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1; -x_167 = l___private_Init_GetElem_0__outOfBounds___rarg(x_166); -x_168 = lean_ctor_get(x_167, 1); -lean_inc(x_168); -lean_dec(x_167); -x_169 = lean_nat_sub(x_165, x_168); -lean_dec(x_168); -lean_dec(x_165); -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_170 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_169, x_134, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_170) == 0) +lean_object* x_62; lean_object* x_63; uint8_t x_64; +x_62 = lean_ctor_get(x_50, 0); +lean_inc(x_62); +x_63 = l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1; +x_64 = lean_nat_dec_eq(x_63, x_62); +lean_dec(x_62); +if (x_64 == 0) { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_171 = lean_ctor_get(x_170, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_170, 1); -lean_inc(x_172); -lean_dec(x_170); -x_173 = lean_box(0); -x_174 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_171, x_173, x_6, x_7, x_8, x_9, x_10, x_11, x_172); +lean_object* x_65; +lean_dec(x_21); +lean_dec(x_20); +lean_dec(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); -return x_174; +lean_dec(x_5); +lean_dec(x_1); +x_65 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_65, 0, x_50); +lean_ctor_set(x_65, 1, x_51); +return x_65; } else { -uint8_t x_175; -lean_dec(x_36); -lean_dec(x_34); -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_175 = !lean_is_exclusive(x_170); -if (x_175 == 0) -{ -return x_170; +lean_object* x_66; +lean_dec(x_50); +x_66 = lean_box(0); +x_26 = x_66; +x_27 = x_51; +goto block_32; } -else -{ -lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_176 = lean_ctor_get(x_170, 0); -x_177 = lean_ctor_get(x_170, 1); -lean_inc(x_177); -lean_inc(x_176); -lean_dec(x_170); -x_178 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_178, 0, x_176); -lean_ctor_set(x_178, 1, x_177); -return x_178; } } } -else +} +} +block_32: { -lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_179 = lean_array_fget(x_35, x_132); -lean_dec(x_35); -x_180 = lean_ctor_get(x_179, 1); -lean_inc(x_180); -lean_dec(x_179); -x_181 = lean_nat_sub(x_165, x_180); -lean_dec(x_180); -lean_dec(x_165); -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_182 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___spec__6(x_181, x_134, x_6, x_7, x_8, x_9, x_10, x_11, x_33); -if (lean_obj_tag(x_182) == 0) +if (lean_obj_tag(x_26) == 0) { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_183 = lean_ctor_get(x_182, 0); -lean_inc(x_183); -x_184 = lean_ctor_get(x_182, 1); -lean_inc(x_184); -lean_dec(x_182); -x_185 = lean_box(0); -x_186 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__5(x_34, x_36, x_183, x_185, x_6, x_7, x_8, x_9, x_10, x_11, x_184); -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_186; +lean_object* x_28; lean_object* x_29; +lean_dec(x_19); +x_28 = lean_box(0); +x_29 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__8(x_21, x_20, x_5, x_1, x_4, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_27); +return x_29; } else { -uint8_t x_187; -lean_dec(x_36); -lean_dec(x_34); +lean_object* x_30; lean_object* x_31; +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); -x_187 = !lean_is_exclusive(x_182); -if (x_187 == 0) -{ -return x_182; -} -else -{ -lean_object* x_188; lean_object* x_189; lean_object* x_190; -x_188 = lean_ctor_get(x_182, 0); -x_189 = lean_ctor_get(x_182, 1); -lean_inc(x_189); -lean_inc(x_188); -lean_dec(x_182); -x_190 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_190, 0, x_188); -lean_ctor_set(x_190, 1, x_189); -return x_190; -} -} -} +lean_dec(x_5); +lean_dec(x_1); +x_30 = lean_ctor_get(x_26, 0); +lean_inc(x_30); +lean_dec(x_26); +if (lean_is_scalar(x_19)) { + x_31 = lean_alloc_ctor(0, 2, 0); +} else { + x_31 = x_19; } +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_27); +return x_31; } } } @@ -18991,7 +18275,7 @@ return x_190; } else { -uint8_t x_195; +uint8_t x_77; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -18999,29 +18283,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); -x_195 = !lean_is_exclusive(x_17); -if (x_195 == 0) +lean_dec(x_1); +x_77 = !lean_is_exclusive(x_16); +if (x_77 == 0) { -return x_17; +return x_16; } else { -lean_object* x_196; lean_object* x_197; lean_object* x_198; -x_196 = lean_ctor_get(x_17, 0); -x_197 = lean_ctor_get(x_17, 1); -lean_inc(x_197); -lean_inc(x_196); -lean_dec(x_17); -x_198 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); -return x_198; +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_16, 0); +x_79 = lean_ctor_get(x_16, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_16); +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; } } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -19031,7 +18315,7 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { _start: { lean_object* x_9; uint8_t x_10; @@ -19057,7 +18341,7 @@ else { uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; x_13 = l_instInhabitedBool; -x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1; +x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1; x_15 = lean_box(x_13); x_16 = l_Lean_PrettyPrinter_Delaborator_withMDatasOptions___rarg(x_15, x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_16) == 0) @@ -19153,7 +18437,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___boxed), 8, 0); return x_1; } } @@ -19391,7 +18675,7 @@ if (x_14 == 0) { lean_object* x_16; lean_object* x_17; x_16 = lean_box(0); -x_17 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6(x_4, x_2, x_12, x_3, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +x_17 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9(x_2, x_3, x_4, x_12, x_16, x_5, x_6, x_7, x_8, x_9, x_10, x_15); return x_17; } else @@ -19412,7 +18696,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6(x_4, x_2, x_12, x_3, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_20); +x_21 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9(x_2, x_3, x_4, x_12, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_20); return x_21; } else @@ -19451,7 +18735,29 @@ return x_25; } } } -LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6(x_1, x_5, x_6, x_4); +lean_dec(x_1); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7___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; @@ -19459,87 +18765,111 @@ x_4 = lean_unbox_usize(x_2); lean_dec(x_2); x_5 = lean_unbox_usize(x_3); lean_dec(x_3); -x_6 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__5(x_1, x_4, x_5); +x_6 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__7(x_1, x_4, x_5); lean_dec(x_1); x_7 = lean_box(x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_3); +lean_object* x_10; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_1); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9___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_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_10); -lean_dec(x_9); +x_6 = l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__9(x_1, x_4, x_5); +lean_dec(x_1); +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_dec(x_2); -return x_13; +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_1); +return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_17; lean_object* x_18; -x_17 = lean_unbox(x_4); -lean_dec(x_4); -x_18 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(x_1, x_2, x_3, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -lean_dec(x_9); +lean_object* x_11; +x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_2); lean_dec(x_1); -return x_18; +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3___boxed(lean_object* x_1) { _start: { -lean_object* x_11; -x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +uint8_t x_2; lean_object* x_3; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__3(x_1); lean_dec(x_1); -return x_11; +x_3 = lean_box(x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_delabAppImplicitCore___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = lean_unbox(x_3); -lean_dec(x_3); -x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_1, x_2, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +lean_object* x_11; +x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -return x_13; +lean_dec(x_3); +lean_dec(x_1); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { _start: { -uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_4); -lean_dec(x_4); -x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5(x_1, x_2, x_3, x_13, 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_5); +lean_dec(x_5); +x_15 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__8(x_1, x_2, x_3, x_4, x_14, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { uint8_t x_13; lean_object* x_14; -x_13 = lean_unbox(x_3); -lean_dec(x_3); -x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6(x_1, x_2, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_13 = lean_unbox(x_4); +lean_dec(x_4); +x_14 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9(x_1, x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___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) { _start: { lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); return x_9; } @@ -20926,14 +20256,6 @@ lean_dec(x_1); return x_10; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_params___default() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -return x_1; -} -} static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default() { _start: { @@ -20958,7 +20280,7 @@ x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_hNames_x3f___default() { _start: { lean_object* x_1; @@ -20966,7 +20288,7 @@ x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_alts___default() { _start: { lean_object* x_1; @@ -20974,7 +20296,7 @@ x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_moreArgs___default() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default() { _start: { lean_object* x_1; @@ -20982,314 +20304,1445 @@ x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___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_object* x_10) { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default() { _start: { -uint8_t x_11; -x_11 = !lean_is_exclusive(x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_4, 3); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); -lean_dec(x_13); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set(x_4, 3, x_15); -x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_16; -} -else -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; -x_17 = lean_ctor_get(x_4, 0); -x_18 = lean_ctor_get(x_4, 1); -x_19 = lean_ctor_get(x_4, 2); -x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); -x_21 = lean_ctor_get(x_4, 3); -x_22 = lean_ctor_get(x_4, 4); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_4); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); -lean_dec(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_1); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_26, 0, x_17); -lean_ctor_set(x_26, 1, x_18); -lean_ctor_set(x_26, 2, x_19); -lean_ctor_set(x_26, 3, x_25); -lean_ctor_set(x_26, 4, x_22); -lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); -x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); -return x_27; -} +lean_object* x_1; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__3___rarg), 10, 0); -return x_2; +lean_object* x_11; +x_11 = lean_apply_9(x_1, x_4, x_5, x_2, x_3, x_6, x_7, x_8, x_9, x_10); +return x_11; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2___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_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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) { _start: { -lean_object* x_12; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_12 = lean_apply_8(x_1, 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_10; lean_object* x_11; +x_10 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg___lambda__1), 10, 3); +lean_closure_set(x_10, 0, x_2); +lean_closure_set(x_10, 1, x_3); +lean_closure_set(x_10, 2, x_4); +x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp___rarg(x_1, x_10, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(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; lean_object* x_19; -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) +{ +return x_11; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_11, 1); lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Expr_bindingBody_x21(x_2); -lean_dec(x_2); -x_16 = lean_expr_instantiate1(x_15, x_4); -lean_dec(x_4); -lean_dec(x_15); -x_17 = lean_apply_1(x_3, x_13); -x_18 = lean_unsigned_to_nat(1u); -x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__3___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_19; +lean_inc(x_13); +lean_dec(x_11); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; +} } else { -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_11); +if (x_16 == 0) { -return x_12; +return x_11; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 0); -x_22 = lean_ctor_get(x_12, 1); -lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_12); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_11, 0); +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_11); +x_19 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_19, 0, x_17); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___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) { -_start: -{ -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; -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_14 = l_Lean_Expr_binderInfo(x_12); -x_15 = l_Lean_Expr_bindingDomain_x21(x_12); -x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2___rarg___lambda__1), 11, 3); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_12); -lean_closure_set(x_16, 2, x_3); -x_17 = 0; -x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_18; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1(lean_object* x_1) { _start: { -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2___rarg), 10, 0); -return x_3; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg), 9, 0); +return x_2; } } -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* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1() { _start: { -lean_object* x_9; lean_object* x_10; -x_9 = lean_box(0); -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_10, 1, x_8); -return x_10; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Unit", 4); +return x_1; } } -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__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) { +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2() { _start: { -lean_object* x_10; -x_10 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, 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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1() { +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__1___boxed), 8, 0); +x_1 = lean_mk_string_from_bytes("unit", 4); return x_1; } } -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* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4() { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__2___boxed), 9, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1; -x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; +x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__3; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_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(lean_object* x_1) { +static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5() { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg), 9, 0); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4; +x_3 = l_Lean_Expr_const___override(x_2, x_1); +return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; uint8_t x_12; +lean_object* x_11; lean_object* x_12; uint8_t x_13; x_11 = lean_array_get_size(x_2); -x_12 = lean_nat_dec_lt(x_1, x_11); +x_12 = lean_unsigned_to_nat(0u); +x_13 = lean_nat_dec_lt(x_12, x_11); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = l_Lean_instInhabitedExpr; +x_15 = l___private_Init_GetElem_0__outOfBounds___rarg(x_14); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_16 = lean_infer_type(x_15, x_6, x_7, x_8, x_9, x_10); +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_11, x_19); lean_dec(x_11); -if (x_12 == 0) +if (x_20 == 0) { -lean_object* x_13; -lean_dec(x_2); -lean_dec(x_1); -x_13 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; +lean_object* x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; +lean_dec(x_17); +lean_inc(x_2); +x_21 = l_Lean_mkAppN(x_1, x_2); +x_22 = 0; +x_23 = 1; +x_24 = 1; +x_25 = l_Lean_Meta_mkLambdaFVars(x_2, x_21, x_22, x_23, x_24, x_6, x_7, x_8, x_9, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_25; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_array_fget(x_2, x_1); -x_15 = lean_unsigned_to_nat(1u); -x_16 = lean_nat_add(x_1, x_15); -lean_dec(x_1); -x_17 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg), 10, 3); -lean_closure_set(x_17, 0, x_16); -lean_closure_set(x_17, 1, x_2); -lean_closure_set(x_17, 2, x_3); -x_18 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg(x_14, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_18; -} -} +lean_object* x_26; uint8_t x_27; +x_26 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; +x_27 = l_Lean_Expr_isConstOf(x_17, x_26); +lean_dec(x_17); +if (x_27 == 0) +{ +lean_object* x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; +lean_inc(x_2); +x_28 = l_Lean_mkAppN(x_1, x_2); +x_29 = 0; +x_30 = 1; +x_31 = 1; +x_32 = l_Lean_Meta_mkLambdaFVars(x_2, x_28, x_29, x_30, x_31, x_6, x_7, x_8, x_9, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_32; } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg), 10, 0); -return x_2; +lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; +x_33 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5; +x_34 = l_Lean_Expr_app___override(x_1, x_33); +x_35 = 0; +x_36 = 1; +x_37 = 1; +x_38 = l_Lean_Meta_mkLambdaFVars(x_2, x_34, x_35, x_36, x_37, x_6, x_7, x_8, x_9, x_18); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_38; } } -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___boxed(lean_object* x_1, lean_object* x_2, 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_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +uint8_t x_39; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -return x_9; -} +x_39 = !lean_is_exclusive(x_16); +if (x_39 == 0) +{ +return x_16; } -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__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -lean_object* x_10; -x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_2); -return x_10; +lean_object* x_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_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNames___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: +} +else { -lean_object* x_10; lean_object* x_11; -x_10 = lean_unsigned_to_nat(0u); -x_11 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(x_10, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; +lean_object* x_43; lean_object* x_44; +x_43 = lean_array_fget(x_2, x_12); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_44 = lean_infer_type(x_43, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_44) == 0) +{ +lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; +x_45 = lean_ctor_get(x_44, 0); +lean_inc(x_45); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_unsigned_to_nat(1u); +x_48 = lean_nat_dec_eq(x_11, x_47); +lean_dec(x_11); +if (x_48 == 0) +{ +lean_object* x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; +lean_dec(x_45); +lean_inc(x_2); +x_49 = l_Lean_mkAppN(x_1, x_2); +x_50 = 0; +x_51 = 1; +x_52 = 1; +x_53 = l_Lean_Meta_mkLambdaFVars(x_2, x_49, x_50, x_51, x_52, x_6, x_7, x_8, x_9, x_46); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_53; } +else +{ +lean_object* x_54; uint8_t x_55; +x_54 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; +x_55 = l_Lean_Expr_isConstOf(x_45, x_54); +lean_dec(x_45); +if (x_55 == 0) +{ +lean_object* x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; lean_object* x_60; +lean_inc(x_2); +x_56 = l_Lean_mkAppN(x_1, x_2); +x_57 = 0; +x_58 = 1; +x_59 = 1; +x_60 = l_Lean_Meta_mkLambdaFVars(x_2, x_56, x_57, x_58, x_59, x_6, x_7, x_8, x_9, x_46); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_60; } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNames(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNames___rarg), 9, 0); -return x_2; +lean_object* x_61; lean_object* x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; lean_object* x_66; +x_61 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5; +x_62 = l_Lean_Expr_app___override(x_1, x_61); +x_63 = 0; +x_64 = 1; +x_65 = 1; +x_66 = l_Lean_Meta_mkLambdaFVars(x_2, x_62, x_63, x_64, x_65, x_6, x_7, x_8, x_9, x_46); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +return x_66; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_4); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_4, 3); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); -lean_dec(x_13); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set(x_4, 3, x_15); -x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; -x_17 = lean_ctor_get(x_4, 0); -x_18 = lean_ctor_get(x_4, 1); +uint8_t x_67; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_2); +lean_dec(x_1); +x_67 = !lean_is_exclusive(x_44); +if (x_67 == 0) +{ +return x_44; +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_44, 0); +x_69 = lean_ctor_get(x_44, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_44); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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) { +_start: +{ +uint8_t x_13; +x_13 = lean_nat_dec_lt(x_4, x_1); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_14 = lean_apply_8(x_3, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); +x_18 = l_Lean_Expr_isLambda(x_16); +if (x_18 == 0) +{ +lean_object* x_19; +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_16); +x_19 = lean_infer_type(x_16, x_8, x_9, x_10, x_11, 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; +x_20 = lean_ctor_get(x_19, 0); +lean_inc(x_20); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___boxed), 10, 1); +lean_closure_set(x_22, 0, 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_23 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg(x_20, x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_21); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_ctor_get(x_6, 3); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_23, 1); +lean_inc(x_26); +lean_dec(x_23); +x_27 = !lean_is_exclusive(x_6); +if (x_27 == 0) +{ +lean_object* x_28; uint8_t x_29; +x_28 = lean_ctor_get(x_6, 3); +lean_dec(x_28); +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_24, 0); +lean_dec(x_30); +lean_ctor_set(x_24, 0, x_25); +x_31 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +return x_31; +} +else +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 1); +lean_inc(x_32); +lean_dec(x_24); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_25); +lean_ctor_set(x_33, 1, x_32); +lean_ctor_set(x_6, 3, x_33); +x_34 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_26); +return x_34; +} +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_35 = lean_ctor_get(x_6, 0); +x_36 = lean_ctor_get(x_6, 1); +x_37 = lean_ctor_get(x_6, 2); +x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*5); +x_39 = lean_ctor_get(x_6, 4); +lean_inc(x_39); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_6); +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_24); + x_41 = lean_box(0); +} +if (lean_is_scalar(x_41)) { + x_42 = lean_alloc_ctor(0, 2, 0); +} else { + x_42 = x_41; +} +lean_ctor_set(x_42, 0, x_25); +lean_ctor_set(x_42, 1, x_40); +x_43 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_43, 0, x_35); +lean_ctor_set(x_43, 1, x_36); +lean_ctor_set(x_43, 2, x_37); +lean_ctor_set(x_43, 3, x_42); +lean_ctor_set(x_43, 4, x_39); +lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_38); +x_44 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_2, x_3, x_4, x_5, x_43, x_7, x_8, x_9, x_10, x_11, x_26); +return x_44; +} +} +else +{ +uint8_t x_45; +lean_dec(x_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_45 = !lean_is_exclusive(x_23); +if (x_45 == 0) +{ +return x_23; +} +else +{ +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_23, 0); +x_47 = lean_ctor_get(x_23, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_23); +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_49; +lean_dec(x_16); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_49 = !lean_is_exclusive(x_19); +if (x_49 == 0) +{ +return x_19; +} +else +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_19, 0); +x_51 = lean_ctor_get(x_19, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_19); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} +} +} +else +{ +lean_object* x_53; +lean_dec(x_16); +x_53 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_17); +return x_53; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 12, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 3); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_4, 3, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_4, 2); +x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); +x_21 = lean_ctor_get(x_4, 3); +x_22 = lean_ctor_get(x_4, 4); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_4); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_18); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_26, 4, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); +x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___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_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_12 = lean_apply_8(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +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_Expr_bindingBody_x21(x_2); +lean_dec(x_2); +x_16 = lean_expr_instantiate1(x_15, x_4); +lean_dec(x_4); +lean_dec(x_15); +x_17 = lean_apply_1(x_3, x_13); +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_12); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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) { +_start: +{ +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; +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_14 = l_Lean_Expr_binderInfo(x_12); +x_15 = l_Lean_Expr_bindingDomain_x21(x_12); +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg___lambda__1), 11, 3); +lean_closure_set(x_16, 0, x_2); +lean_closure_set(x_16, 1, x_12); +lean_closure_set(x_16, 2, x_3); +x_17 = 0; +x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg), 10, 0); +return x_3; +} +} +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* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +return x_10; +} +} +LEAN_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__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; +x_10 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_10; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__1___boxed), 8, 0); +return x_1; +} +} +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(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; +x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +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(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__6___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 3); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_4, 3, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_4, 2); +x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); +x_21 = lean_ctor_get(x_4, 3); +x_22 = lean_ctor_get(x_4, 4); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_4); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_18); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_26, 4, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); +x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__6(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__6___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___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_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_12 = lean_apply_8(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +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_Expr_bindingBody_x21(x_2); +lean_dec(x_2); +x_16 = lean_expr_instantiate1(x_15, x_4); +lean_dec(x_4); +lean_dec(x_15); +x_17 = lean_apply_1(x_3, x_13); +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__6___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_12); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +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; +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_14 = l_Lean_Expr_binderInfo(x_12); +x_15 = l_Lean_Expr_bindingDomain_x21(x_12); +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___rarg___lambda__1), 11, 3); +lean_closure_set(x_16, 0, x_2); +lean_closure_set(x_16, 1, x_12); +lean_closure_set(x_16, 2, x_3); +x_17 = 0; +x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___rarg), 10, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__4___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, 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 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; +x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__5___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__4(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__4___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 3); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_4, 3, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_4, 2); +x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); +x_21 = lean_ctor_get(x_4, 3); +x_22 = lean_ctor_get(x_4, 4); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_dec(x_4); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_18); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_26, 4, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); +x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); +return x_27; +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___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_object* x_11) { +_start: +{ +lean_object* x_12; +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_12 = lean_apply_8(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +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_Expr_bindingBody_x21(x_2); +lean_dec(x_2); +x_16 = lean_expr_instantiate1(x_15, x_4); +lean_dec(x_4); +lean_dec(x_15); +x_17 = lean_apply_1(x_3, x_13); +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__9___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_19; +} +else +{ +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) +{ +return x_12; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); +lean_inc(x_22); +lean_inc(x_21); +lean_dec(x_12); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +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_14 = l_Lean_Expr_binderInfo(x_12); +x_15 = l_Lean_Expr_bindingDomain_x21(x_12); +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg___lambda__1), 11, 3); +lean_closure_set(x_16, 0, x_2); +lean_closure_set(x_16, 1, x_12); +lean_closure_set(x_16, 2, x_3); +x_17 = 0; +x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg), 10, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__7___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; +x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__8___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__7(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__7___rarg), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_object* x_11, lean_object* x_12, lean_object* x_13) { +_start: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_unsigned_to_nat(1u); +x_15 = lean_nat_add(x_1, x_14); +lean_dec(x_1); +x_16 = l_Lean_Syntax_getId(x_6); +lean_dec(x_6); +x_17 = lean_array_push(x_2, x_16); +x_18 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(x_3, x_4, x_5, x_15, x_17, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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) { +_start: +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_13 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(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_ctor_get(x_13, 1); +lean_inc(x_15); +lean_dec(x_13); +x_16 = l_Lean_Expr_bindingName_x21(x_14); +lean_dec(x_14); +x_17 = lean_erase_macro_scopes(x_16); +x_18 = lean_array_get_size(x_2); +x_19 = lean_nat_sub(x_1, x_18); +x_20 = lean_nat_dec_le(x_19, x_4); +lean_dec(x_19); +if (x_20 == 0) +{ +uint8_t x_21; +lean_dec(x_18); +x_21 = l_Array_contains___at_Lean_findField_x3f___spec__1(x_5, x_17); +if (x_21 == 0) +{ +lean_object* x_22; lean_object* x_23; +lean_dec(x_17); +x_22 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg___lambda__1), 13, 5); +lean_closure_set(x_22, 0, x_4); +lean_closure_set(x_22, 1, x_5); +lean_closure_set(x_22, 2, x_1); +lean_closure_set(x_22, 3, x_2); +lean_closure_set(x_22, 4, x_3); +x_23 = l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_24 = lean_ctor_get(x_8, 1); +lean_inc(x_24); +x_25 = l_Lean_LocalContext_getUnusedName(x_24, x_17); +x_26 = lean_unsigned_to_nat(1u); +x_27 = lean_nat_add(x_4, x_26); +lean_dec(x_4); +lean_inc(x_25); +x_28 = lean_array_push(x_5, x_25); +x_29 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 12, 5); +lean_closure_set(x_29, 0, x_1); +lean_closure_set(x_29, 1, x_2); +lean_closure_set(x_29, 2, x_3); +lean_closure_set(x_29, 3, x_27); +lean_closure_set(x_29, 4, x_28); +x_30 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg(x_25, x_29, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_35; +lean_dec(x_17); +x_31 = lean_nat_add(x_4, x_18); +x_32 = lean_nat_sub(x_31, x_1); +lean_dec(x_31); +x_33 = lean_nat_dec_lt(x_32, x_18); +lean_dec(x_18); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_4, x_34); +lean_dec(x_4); +if (x_33 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; +lean_dec(x_32); +x_36 = l_Lean_instInhabitedName; +x_37 = l___private_Init_GetElem_0__outOfBounds___rarg(x_36); +lean_inc(x_37); +x_38 = lean_array_push(x_5, x_37); +x_39 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 12, 5); +lean_closure_set(x_39, 0, x_1); +lean_closure_set(x_39, 1, x_2); +lean_closure_set(x_39, 2, x_3); +lean_closure_set(x_39, 3, x_35); +lean_closure_set(x_39, 4, x_38); +x_40 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__4___rarg(x_37, x_39, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_40; +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_41 = lean_array_fget(x_2, x_32); +lean_dec(x_32); +lean_inc(x_41); +x_42 = lean_array_push(x_5, x_41); +x_43 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 12, 5); +lean_closure_set(x_43, 0, x_1); +lean_closure_set(x_43, 1, x_2); +lean_closure_set(x_43, 2, x_3); +lean_closure_set(x_43, 3, x_35); +lean_closure_set(x_43, 4, x_42); +x_44 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__7___rarg(x_41, x_43, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +return x_44; +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg), 12, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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) { +_start: +{ +lean_object* x_11; +x_11 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___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); +lean_dec(x_1); +return x_9; +} +} +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__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_10; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_2); +return x_10; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_unsigned_to_nat(0u); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_13 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(x_1, x_2, x_3, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; +} +} +LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg), 10, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___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_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_dec(x_6); +x_14 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_14, 0, x_1); +x_15 = lean_array_push(x_2, x_14); +x_16 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg(x_3, x_4, x_5, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +_start: +{ +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_array_get_size(x_4); +x_13 = lean_array_get_size(x_1); +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_2); +lean_dec(x_1); +x_15 = lean_apply_8(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +return x_15; +} +else +{ +lean_object* x_16; +x_16 = lean_array_fget(x_1, x_12); +lean_dec(x_12); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_box(0); +x_18 = lean_array_push(x_4, x_17); +x_4 = x_18; +goto _start; +} +else +{ +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_16, 0); +lean_inc(x_20); +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_2); +x_21 = l_Lean_PrettyPrinter_Delaborator_getUnusedName(x_20, x_2, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_21) == 0) +{ +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; +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); +lean_inc(x_22); +x_24 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg___lambda__1), 13, 5); +lean_closure_set(x_24, 0, x_22); +lean_closure_set(x_24, 1, x_4); +lean_closure_set(x_24, 2, x_1); +lean_closure_set(x_24, 3, x_2); +lean_closure_set(x_24, 4, x_3); +x_25 = 0; +x_26 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1; +x_27 = 1; +x_28 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_22, x_25, x_26, x_24, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +return x_28; +} +else +{ +uint8_t x_29; +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_29 = !lean_is_exclusive(x_21); +if (x_29 == 0) +{ +return x_21; +} +else +{ +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_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg), 11, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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_object* x_10) { +_start: +{ +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 3); +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_4, 3, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; 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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); x_19 = lean_ctor_get(x_4, 2); x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); x_21 = lean_ctor_get(x_4, 3); @@ -21320,87 +21773,49 @@ return x_27; } } } -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* x_1, lean_object* x_2, lean_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_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__3(lean_object* x_1) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Expr_appFn_x21(x_10); -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_14; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__3___rarg), 10, 0); +return x_2; } } -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* x_1, lean_object* x_2, lean_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_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2___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_object* x_11) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); +lean_object* x_12; lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Expr_appArg_x21(x_10); -lean_dec(x_10); -x_13 = lean_unsigned_to_nat(1u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Expr_isApp(x_11); -lean_dec(x_11); -if (x_13 == 0) -{ -lean_object* x_14; -lean_dec(x_2); -x_14 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -return x_14; -} -else -{ -lean_object* x_15; lean_object* x_16; -lean_inc(x_2); -x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__1), 9, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); +lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -lean_inc(x_3); -x_16 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__2(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_16) == 0) +x_12 = lean_apply_8(x_1, 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_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_apply_1(x_2, x_17); -x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__4(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -return x_20; +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_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_Expr_bindingBody_x21(x_2); +lean_dec(x_2); +x_16 = lean_expr_instantiate1(x_15, x_4); +lean_dec(x_4); +lean_dec(x_15); +x_17 = lean_apply_1(x_3, x_13); +x_18 = lean_unsigned_to_nat(1u); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__3___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +return x_19; } else { -uint8_t x_21; +uint8_t x_20; +lean_dec(x_10); +lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -21408,877 +21823,775 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) +x_20 = !lean_is_exclusive(x_12); +if (x_20 == 0) { -return x_16; +return x_12; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_12, 0); +x_22 = lean_ctor_get(x_12, 1); lean_inc(x_22); -lean_dec(x_16); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} +lean_inc(x_21); +lean_dec(x_12); +x_23 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_23, 0, x_21); +lean_ctor_set(x_23, 1, x_22); +return x_23; } } } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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) { _start: { -lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +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; +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_14 = l_Lean_Expr_binderInfo(x_12); +x_15 = l_Lean_Expr_bindingDomain_x21(x_12); +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2___rarg___lambda__1), 11, 3); +lean_closure_set(x_16, 0, x_2); +lean_closure_set(x_16, 1, x_12); +lean_closure_set(x_16, 2, x_3); +x_17 = 0; +x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_18; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2(lean_object* x_1, lean_object* x_2) { +_start: { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2___rarg), 10, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___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) { +_start: { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_array_push(x_1, x_11); -lean_ctor_set(x_9, 0, x_12); -return x_9; +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); +lean_closure_set(x_10, 0, x_2); +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; +x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_12; } -else +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__1(lean_object* x_1) { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_9); -x_15 = lean_array_push(x_1, x_13); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__1___rarg), 9, 0); +return x_2; } } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -uint8_t x_17; -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_9); -if (x_17 == 0) +lean_object* x_11; uint8_t x_12; +x_11 = lean_array_get_size(x_1); +x_12 = lean_nat_dec_lt(x_3, x_11); +lean_dec(x_11); +if (x_12 == 0) { -return x_9; +lean_object* x_13; +lean_dec(x_3); +lean_dec(x_1); +x_13 = lean_apply_7(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_9, 0); -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_9); -x_20 = lean_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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_array_fget(x_1, x_3); +x_15 = lean_unsigned_to_nat(1u); +x_16 = lean_nat_add(x_3, x_15); +lean_dec(x_3); +x_17 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___rarg), 10, 3); +lean_closure_set(x_17, 0, x_1); +lean_closure_set(x_17, 1, x_2); +lean_closure_set(x_17, 2, x_16); +x_18 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___spec__1___rarg(x_14, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_18; } } } -static lean_object* _init_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_delabAppMatch_usingNames(lean_object* x_1) { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg___boxed), 8, 1); -lean_closure_set(x_2, 0, x_1); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___rarg), 10, 0); return x_2; } } -static lean_object* _init_l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___lambda__1), 8, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_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; -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); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -x_20 = lean_infer_type(x_19, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_20) == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; uint8_t x_35; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_1, 5); -x_24 = lean_array_get_size(x_23); -x_25 = lean_nat_dec_lt(x_5, x_24); -lean_dec(x_24); -x_26 = l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__1; -x_27 = l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2; -x_28 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__1), 9, 2); -lean_closure_set(x_28, 0, x_26); -lean_closure_set(x_28, 1, x_27); -x_29 = lean_ctor_get(x_8, 3); -lean_inc(x_29); -x_30 = lean_ctor_get(x_8, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_8, 1); -lean_inc(x_31); -x_32 = lean_ctor_get(x_8, 2); -lean_inc(x_32); -x_33 = lean_ctor_get_uint8(x_8, sizeof(void*)*5); -x_34 = lean_ctor_get(x_8, 4); -lean_inc(x_34); -x_35 = !lean_is_exclusive(x_29); -if (x_35 == 0) -{ -lean_object* x_36; lean_object* x_37; -x_36 = lean_ctor_get(x_29, 0); -lean_dec(x_36); -lean_ctor_set(x_29, 0, x_21); -x_37 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_37, 0, x_30); -lean_ctor_set(x_37, 1, x_31); -lean_ctor_set(x_37, 2, x_32); -lean_ctor_set(x_37, 3, x_29); -lean_ctor_set(x_37, 4, x_34); -lean_ctor_set_uint8(x_37, sizeof(void*)*5, x_33); -if (x_25 == 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_38; lean_object* x_39; lean_object* x_40; -x_38 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_39 = l___private_Init_GetElem_0__outOfBounds___rarg(x_38); -lean_inc(x_13); +lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_40 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(x_15, x_39, x_28, x_37, x_9, x_10, x_11, x_12, x_13, x_22); -if (lean_obj_tag(x_40) == 0) -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = lean_nat_add(x_5, x_17); -lean_dec(x_5); -x_44 = lean_array_push(x_7, x_41); -x_4 = x_18; -x_5 = x_43; -x_6 = lean_box(0); -x_7 = x_44; -x_14 = x_42; -goto _start; -} -else -{ -uint8_t x_46; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_12); lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_46 = !lean_is_exclusive(x_40); -if (x_46 == 0) -{ -return x_40; +x_13 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_12, x_1); +lean_ctor_set(x_9, 0, x_13); +return x_9; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_40, 0); -x_48 = lean_ctor_get(x_40, 1); -lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_40); -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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = lean_ctor_get(x_9, 0); +x_15 = lean_ctor_get(x_9, 1); +lean_inc(x_15); +lean_inc(x_14); +lean_dec(x_9); +x_16 = lean_ctor_get(x_14, 0); +lean_inc(x_16); +lean_dec(x_14); +x_17 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_16, x_1); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_15); +return x_18; } } } -else +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_50; lean_object* x_51; -x_50 = lean_array_fget(x_23, x_5); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_51 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(x_15, x_50, x_28, x_37, x_9, x_10, x_11, x_12, x_13, x_22); -if (lean_obj_tag(x_51) == 0) +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_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); -lean_inc(x_53); -lean_dec(x_51); -x_54 = lean_nat_add(x_5, x_17); -lean_dec(x_5); -x_55 = lean_array_push(x_7, x_52); -x_4 = x_18; -x_5 = x_54; -x_6 = lean_box(0); -x_7 = x_55; -x_14 = x_53; -goto _start; +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 { -uint8_t x_57; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); +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_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_57 = !lean_is_exclusive(x_51); -if (x_57 == 0) +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; +} +} +} +static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1() { +_start: { -return x_51; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unknown constant '", 18); +return x_1; } -else +} +static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2() { +_start: { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_51, 0); -x_59 = lean_ctor_get(x_51, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_51); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("'", 1); +return x_1; } } -else +static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4() { +_start: { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_29, 1); -lean_inc(x_61); -lean_dec(x_29); -x_62 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_62, 0, x_21); -lean_ctor_set(x_62, 1, x_61); -x_63 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_63, 0, x_30); -lean_ctor_set(x_63, 1, x_31); -lean_ctor_set(x_63, 2, x_32); -lean_ctor_set(x_63, 3, x_62); -lean_ctor_set(x_63, 4, x_34); -lean_ctor_set_uint8(x_63, sizeof(void*)*5, x_33); -if (x_25 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_65 = l___private_Init_GetElem_0__outOfBounds___rarg(x_64); +lean_object* x_9; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_66 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(x_15, x_65, x_28, x_63, x_9, x_10, x_11, x_12, x_13, x_22); -if (lean_obj_tag(x_66) == 0) +lean_dec(x_11); +lean_inc(x_1); +x_14 = lean_environment_find(x_13, x_1); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); -lean_inc(x_68); -lean_dec(x_66); -x_69 = lean_nat_add(x_5, x_17); -lean_dec(x_5); -x_70 = lean_array_push(x_7, x_67); -x_4 = x_18; -x_5 = x_69; -x_6 = lean_box(0); -x_7 = x_70; -x_14 = x_68; -goto _start; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_9); +x_15 = lean_box(0); +x_16 = l_Lean_Expr_const___override(x_1, x_15); +x_17 = l_Lean_MessageData_ofExpr(x_16); +x_18 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__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_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__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_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +return x_22; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_72 = lean_ctor_get(x_66, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_66, 1); -lean_inc(x_73); -if (lean_is_exclusive(x_66)) { - lean_ctor_release(x_66, 0); - lean_ctor_release(x_66, 1); - x_74 = x_66; -} else { - lean_dec_ref(x_66); - x_74 = lean_box(0); -} -if (lean_is_scalar(x_74)) { - x_75 = lean_alloc_ctor(1, 2, 0); -} else { - x_75 = x_74; -} -lean_ctor_set(x_75, 0, x_72); -lean_ctor_set(x_75, 1, x_73); -return x_75; +lean_object* x_23; +lean_dec(x_1); +x_23 = lean_ctor_get(x_14, 0); +lean_inc(x_23); +lean_dec(x_14); +lean_ctor_set(x_9, 0, x_23); +return x_9; } } else { -lean_object* x_76; lean_object* x_77; -x_76 = lean_array_fget(x_23, x_5); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); -lean_inc(x_9); -x_77 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___rarg(x_15, x_76, x_28, x_63, x_9, x_10, x_11, x_12, x_13, x_22); -if (lean_obj_tag(x_77) == 0) +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +lean_inc(x_1); +x_27 = lean_environment_find(x_26, x_1); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_77, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_77, 1); -lean_inc(x_79); -lean_dec(x_77); -x_80 = lean_nat_add(x_5, x_17); -lean_dec(x_5); -x_81 = lean_array_push(x_7, x_78); -x_4 = x_18; -x_5 = x_80; -x_6 = lean_box(0); -x_7 = x_81; -x_14 = x_79; -goto _start; +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_28 = lean_box(0); +x_29 = l_Lean_Expr_const___override(x_1, x_28); +x_30 = l_Lean_MessageData_ofExpr(x_29); +x_31 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_25); +return x_35; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_83 = lean_ctor_get(x_77, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_77, 1); -lean_inc(x_84); -if (lean_is_exclusive(x_77)) { - lean_ctor_release(x_77, 0); - lean_ctor_release(x_77, 1); - x_85 = x_77; -} else { - lean_dec_ref(x_77); - x_85 = lean_box(0); -} -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_object* x_36; lean_object* x_37; +lean_dec(x_1); +x_36 = lean_ctor_get(x_27, 0); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_25); +return x_37; } } } } -else +LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___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, lean_object* x_10) { +_start: { -uint8_t x_87; -lean_dec(x_18); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_5); -x_87 = !lean_is_exclusive(x_20); -if (x_87 == 0) +lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg___lambda__1), 10, 3); +lean_closure_set(x_11, 0, x_2); +lean_closure_set(x_11, 1, x_4); +lean_closure_set(x_11, 2, x_5); +x_12 = 0; +x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(x_1, x_12, x_11, x_3, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_13) == 0) { -return x_20; +uint8_t x_14; +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) +{ +return x_13; } else { -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_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_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; } } +else +{ +uint8_t x_18; +x_18 = !lean_is_exclusive(x_13); +if (x_18 == 0) +{ +return x_13; } else { -lean_object* x_91; +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); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_5); -lean_dec(x_4); -x_91 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_91, 0, x_7); -lean_ctor_set(x_91, 1, x_14); -return x_91; +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_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___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_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4(lean_object* x_1) { _start: { -lean_object* x_11; -x_11 = lean_apply_9(x_1, x_4, x_5, x_2, x_3, x_6, x_7, x_8, x_9, x_10); -return x_11; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___rarg___boxed), 10, 0); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___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_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; -x_10 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg___lambda__1), 10, 3); -lean_closure_set(x_10, 0, x_2); -lean_closure_set(x_10, 1, x_3); -lean_closure_set(x_10, 2, x_4); -x_11 = l___private_Lean_Meta_Basic_0__Lean_Meta_forallTelescopeReducingImp___rarg(x_1, x_10, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) -{ -uint8_t x_12; -x_12 = !lean_is_exclusive(x_11); -if (x_12 == 0) -{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabFVar___spec__1___closed__1; +x_10 = lean_panic_fn(x_9, x_1); +x_11 = lean_apply_7(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_11; } -else +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_11, 0); -x_14 = lean_ctor_get(x_11, 1); -lean_inc(x_14); +uint8_t x_11; +x_11 = !lean_is_exclusive(x_4); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_4, 3); +x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); -lean_dec(x_11); +lean_dec(x_12); +x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +lean_dec(x_13); x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 0, x_1); lean_ctor_set(x_15, 1, x_14); -return x_15; -} -} -else -{ -uint8_t x_16; -x_16 = !lean_is_exclusive(x_11); -if (x_16 == 0) -{ -return x_11; +lean_ctor_set(x_4, 3, x_15); +x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_11, 0); -x_18 = lean_ctor_get(x_11, 1); +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; +x_17 = lean_ctor_get(x_4, 0); +x_18 = lean_ctor_get(x_4, 1); +x_19 = lean_ctor_get(x_4, 2); +x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); +x_21 = lean_ctor_get(x_4, 3); +x_22 = lean_ctor_get(x_4, 4); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -lean_dec(x_11); -x_19 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_18); -return x_19; -} +lean_dec(x_4); +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); +lean_dec(x_23); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_24); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_17); +lean_ctor_set(x_26, 1, x_18); +lean_ctor_set(x_26, 2, x_19); +lean_ctor_set(x_26, 3, x_25); +lean_ctor_set(x_26, 4, x_22); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); +x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); +return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg), 9, 0); -return x_2; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appFn_x21(x_10); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_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_dec(x_3); -x_11 = lean_ctor_get(x_1, 4); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -x_12 = lean_array_get_size(x_11); -lean_dec(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appArg_x21(x_10); +lean_dec(x_10); x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_add(x_13, x_12); -lean_dec(x_12); -x_15 = lean_array_get_size(x_2); -x_16 = l_Array_toSubarray___rarg(x_2, x_14, x_15); -x_17 = l_Array_ofSubarray___rarg(x_16); -x_18 = lean_array_get_size(x_17); -x_19 = lean_mk_empty_array_with_capacity(x_18); -x_20 = lean_unsigned_to_nat(0u); -x_21 = l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5(x_1, x_17, x_17, x_18, x_20, lean_box(0), x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_17); -lean_dec(x_1); -return x_21; +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns(lean_object* x_1, lean_object* x_2, lean_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_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); -x_11 = lean_box(0); -x_12 = !lean_is_exclusive(x_2); -if (x_12 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; +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_14 = lean_unsigned_to_nat(0u); +x_15 = lean_nat_dec_eq(x_1, x_14); +if (x_15 == 0) { -lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_2, 0); -lean_dec(x_13); -x_14 = 1; -lean_ctor_set(x_2, 0, x_11); -lean_ctor_set_uint8(x_2, sizeof(void*)*5, x_14); -x_15 = lean_unsigned_to_nat(0u); +if (lean_obj_tag(x_12) == 5) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_12); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_1, x_16); +lean_dec(x_1); +lean_inc(x_3); +x_18 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__6), 10, 3); +lean_closure_set(x_18, 0, x_17); +lean_closure_set(x_18, 1, x_2); +lean_closure_set(x_18, 2, x_3); +lean_inc(x_9); +lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(x_10, x_15, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_10); -if (lean_obj_tag(x_16) == 0) +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__7(x_18, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___lambda__1), 10, 1); -lean_closure_set(x_19, 0, x_1); -x_20 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg(x_17, x_19, x_2, x_3, x_4, x_5, x_6, x_7, 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_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_apply_1(x_3, x_20); +x_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +return x_23; } else { -uint8_t x_21; -lean_dec(x_2); +uint8_t x_24; +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_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) +x_24 = !lean_is_exclusive(x_19); +if (x_24 == 0) { -return x_16; +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_16); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_19, 0); +x_26 = lean_ctor_get(x_19, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_19); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } else { -lean_object* x_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; -x_25 = lean_ctor_get(x_2, 1); -x_26 = lean_ctor_get(x_2, 2); -x_27 = lean_ctor_get(x_2, 3); -x_28 = lean_ctor_get(x_2, 4); -lean_inc(x_28); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_2); -x_29 = 1; -x_30 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_30, 0, x_11); -lean_ctor_set(x_30, 1, x_25); -lean_ctor_set(x_30, 2, x_26); -lean_ctor_set(x_30, 3, x_27); -lean_ctor_set(x_30, 4, x_28); -lean_ctor_set_uint8(x_30, sizeof(void*)*5, x_29); -x_31 = lean_unsigned_to_nat(0u); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_32 = l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(x_10, x_31, x_9, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_10); -if (lean_obj_tag(x_32) == 0) -{ -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); -x_35 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___lambda__1), 10, 1); -lean_closure_set(x_35, 0, x_1); -x_36 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg(x_33, x_35, x_30, x_3, x_4, x_5, x_6, x_7, x_34); -return x_36; +lean_object* x_28; +lean_dec(x_12); +lean_dec(x_3); +lean_dec(x_1); +x_28 = lean_apply_7(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_28; +} } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -lean_dec(x_30); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); +lean_object* x_29; +lean_dec(x_12); lean_dec(x_3); lean_dec(x_1); -x_37 = lean_ctor_get(x_32, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_32, 1); -lean_inc(x_38); -if (lean_is_exclusive(x_32)) { - lean_ctor_release(x_32, 0); - lean_ctor_release(x_32, 1); - x_39 = x_32; -} else { - lean_dec_ref(x_32); - x_39 = lean_box(0); +x_29 = lean_apply_7(x_2, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +return x_29; } -if (lean_is_scalar(x_39)) { - x_40 = lean_alloc_ctor(1, 2, 0); -} else { - x_40 = x_39; } -lean_ctor_set(x_40, 0, x_37); -lean_ctor_set(x_40, 1, x_38); -return x_40; } +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(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___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { -lean_object* x_15; -x_15 = l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___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); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_15; -} +uint8_t x_5; +x_5 = lean_usize_dec_eq(x_2, x_3); +if (x_5 == 0) +{ +lean_object* x_6; size_t x_7; size_t x_8; +x_6 = lean_array_uget(x_1, x_2); +x_7 = 1; +x_8 = lean_usize_add(x_2, x_7); +if (lean_obj_tag(x_6) == 0) +{ +x_2 = x_8; +goto _start; } -static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Unit", 4); -return x_1; +lean_object* x_10; lean_object* x_11; +x_10 = lean_ctor_get(x_6, 0); +lean_inc(x_10); +lean_dec(x_6); +x_11 = lean_array_push(x_4, x_10); +x_2 = x_8; +x_4 = x_11; +goto _start; } } -static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +return x_4; } } -static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__3() { +} +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unit", 4); -return x_1; +uint8_t x_4; +x_4 = lean_nat_dec_lt(x_2, x_3); +if (x_4 == 0) +{ +lean_object* x_5; +lean_dec(x_3); +lean_dec(x_2); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_5; } +else +{ +lean_object* x_6; uint8_t x_7; +x_6 = lean_array_get_size(x_1); +x_7 = lean_nat_dec_le(x_3, x_6); +lean_dec(x_6); +if (x_7 == 0) +{ +lean_object* x_8; +lean_dec(x_3); +lean_dec(x_2); +x_8 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +return x_8; } -static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1; -x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__3; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; +size_t x_9; size_t x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_usize_of_nat(x_2); +lean_dec(x_2); +x_10 = lean_usize_of_nat(x_3); +lean_dec(x_3); +x_11 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_12 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_1, x_9, x_10, x_11); +return x_12; } } -static lean_object* _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4; -x_3 = l_Lean_Expr_const___override(x_2, x_1); -return x_3; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___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; uint8_t x_13; -x_11 = lean_array_get_size(x_2); -x_12 = lean_unsigned_to_nat(0u); -x_13 = lean_nat_dec_lt(x_12, x_11); -if (x_13 == 0) +lean_object* x_11; +x_11 = l_Lean_PrettyPrinter_Delaborator_delab(x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = l_Lean_instInhabitedExpr; -x_15 = l___private_Init_GetElem_0__outOfBounds___rarg(x_14); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_16 = lean_infer_type(x_15, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_16) == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 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_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; lean_object* x_23; lean_object* x_24; +x_13 = lean_ctor_get(x_11, 0); +x_14 = lean_ctor_get(x_1, 1); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 2); +lean_inc(x_15); +x_16 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_17 = lean_ctor_get(x_1, 3); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +x_18 = lean_ctor_get(x_1, 4); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_unsigned_to_nat(1u); -x_20 = lean_nat_dec_eq(x_11, x_19); -lean_dec(x_11); -if (x_20 == 0) -{ -lean_object* x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; lean_object* x_25; -lean_dec(x_17); -lean_inc(x_2); -x_21 = l_Lean_mkAppN(x_1, x_2); -x_22 = 0; -x_23 = 1; -x_24 = 1; -x_25 = l_Lean_Meta_mkLambdaFVars(x_2, x_21, x_22, x_23, x_24, x_6, x_7, x_8, x_9, x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_25; -} -else -{ -lean_object* x_26; uint8_t x_27; -x_26 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; -x_27 = l_Lean_Expr_isConstOf(x_17, x_26); -lean_dec(x_17); -if (x_27 == 0) -{ -lean_object* x_28; uint8_t x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; -lean_inc(x_2); -x_28 = l_Lean_mkAppN(x_1, x_2); -x_29 = 0; -x_30 = 1; -x_31 = 1; -x_32 = l_Lean_Meta_mkLambdaFVars(x_2, x_28, x_29, x_30, x_31, x_6, x_7, x_8, x_9, x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_32; +x_19 = lean_ctor_get(x_1, 5); +lean_inc(x_19); +x_20 = lean_ctor_get(x_1, 6); +lean_inc(x_20); +x_21 = lean_array_push(x_20, x_3); +x_22 = lean_ctor_get(x_1, 7); +lean_inc(x_22); +lean_dec(x_1); +x_23 = lean_array_push(x_22, x_13); +x_24 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_24, 0, x_2); +lean_ctor_set(x_24, 1, x_14); +lean_ctor_set(x_24, 2, x_15); +lean_ctor_set(x_24, 3, x_17); +lean_ctor_set(x_24, 4, x_18); +lean_ctor_set(x_24, 5, x_19); +lean_ctor_set(x_24, 6, x_21); +lean_ctor_set(x_24, 7, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*8, x_16); +lean_ctor_set(x_11, 0, x_24); +return x_11; } else { -lean_object* x_33; lean_object* x_34; uint8_t x_35; uint8_t x_36; uint8_t x_37; lean_object* x_38; -x_33 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5; -x_34 = l_Lean_Expr_app___override(x_1, x_33); -x_35 = 0; -x_36 = 1; -x_37 = 1; -x_38 = l_Lean_Meta_mkLambdaFVars(x_2, x_34, x_35, x_36, x_37, x_6, x_7, x_8, x_9, x_18); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); +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; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_25 = lean_ctor_get(x_11, 0); +x_26 = lean_ctor_get(x_11, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_11); +x_27 = lean_ctor_get(x_1, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_1, 2); +lean_inc(x_28); +x_29 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_30 = lean_ctor_get(x_1, 3); +lean_inc(x_30); +x_31 = lean_ctor_get(x_1, 4); +lean_inc(x_31); +x_32 = lean_ctor_get(x_1, 5); +lean_inc(x_32); +x_33 = lean_ctor_get(x_1, 6); +lean_inc(x_33); +x_34 = lean_array_push(x_33, x_3); +x_35 = lean_ctor_get(x_1, 7); +lean_inc(x_35); +lean_dec(x_1); +x_36 = lean_array_push(x_35, x_25); +x_37 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_37, 0, x_2); +lean_ctor_set(x_37, 1, x_27); +lean_ctor_set(x_37, 2, x_28); +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_34); +lean_ctor_set(x_37, 7, x_36); +lean_ctor_set_uint8(x_37, sizeof(void*)*8, x_29); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_37); +lean_ctor_set(x_38, 1, x_26); return x_38; } } -} else { uint8_t x_39; -lean_dec(x_11); -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_39 = !lean_is_exclusive(x_16); +x_39 = !lean_is_exclusive(x_11); if (x_39 == 0) { -return x_16; +return x_11; } 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); +x_40 = lean_ctor_get(x_11, 0); +x_41 = lean_ctor_get(x_11, 1); lean_inc(x_41); lean_inc(x_40); -lean_dec(x_16); +lean_dec(x_11); x_42 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); @@ -22286,329 +22599,381 @@ return x_42; } } } -else +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, 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_43; lean_object* x_44; -x_43 = lean_array_fget(x_2, x_12); +uint8_t x_14; +x_14 = lean_nat_dec_le(x_4, x_3); +if (x_14 == 0) +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_nat_dec_eq(x_2, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; +x_17 = lean_unsigned_to_nat(1u); +x_18 = lean_nat_sub(x_2, x_17); +lean_dec(x_2); +x_25 = lean_ctor_get(x_6, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_6, 5); +lean_inc(x_26); +x_27 = lean_ctor_get(x_25, 2); +lean_inc(x_27); +x_28 = lean_array_get_size(x_27); +x_29 = lean_nat_dec_lt(x_3, x_28); +lean_dec(x_28); +x_30 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___lambda__1), 10, 2); +lean_closure_set(x_30, 0, x_6); +lean_closure_set(x_30, 1, x_25); +x_31 = lean_ctor_get(x_7, 0); +x_32 = lean_ctor_get(x_7, 1); +x_33 = lean_ctor_get(x_7, 2); +x_34 = lean_ctor_get_uint8(x_7, sizeof(void*)*5); +x_35 = lean_ctor_get(x_7, 4); +x_36 = lean_array_get_size(x_26); +x_37 = lean_nat_dec_lt(x_3, x_36); +lean_dec(x_36); +if (x_29 == 0) +{ +lean_object* x_38; lean_object* x_39; +lean_dec(x_27); +x_38 = l_instInhabitedNat; +x_39 = l___private_Init_GetElem_0__outOfBounds___rarg(x_38); +if (x_37 == 0) +{ +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_dec(x_26); +x_40 = l_Lean_instInhabitedSubExpr; +x_41 = l___private_Init_GetElem_0__outOfBounds___rarg(x_40); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +x_42 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_42, 0, x_31); +lean_ctor_set(x_42, 1, x_32); +lean_ctor_set(x_42, 2, x_33); +lean_ctor_set(x_42, 3, x_41); +lean_ctor_set(x_42, 4, x_35); +lean_ctor_set_uint8(x_42, sizeof(void*)*5, x_34); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_44 = lean_infer_type(x_43, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_44) == 0) +lean_inc(x_1); +x_43 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(x_39, x_1, x_30, x_42, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_45 = lean_ctor_get(x_44, 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); -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = lean_unsigned_to_nat(1u); -x_48 = lean_nat_dec_eq(x_11, x_47); -lean_dec(x_11); -if (x_48 == 0) -{ -lean_object* x_49; uint8_t x_50; uint8_t x_51; uint8_t x_52; lean_object* x_53; -lean_dec(x_45); -lean_inc(x_2); -x_49 = l_Lean_mkAppN(x_1, x_2); -x_50 = 0; -x_51 = 1; -x_52 = 1; -x_53 = l_Lean_Meta_mkLambdaFVars(x_2, x_49, x_50, x_51, x_52, x_6, x_7, x_8, x_9, x_46); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_53; +lean_dec(x_43); +x_46 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_46, 0, x_44); +x_19 = x_46; +x_20 = x_45; +goto block_24; } else { -lean_object* x_54; uint8_t x_55; -x_54 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2; -x_55 = l_Lean_Expr_isConstOf(x_45, x_54); -lean_dec(x_45); -if (x_55 == 0) -{ -lean_object* x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; lean_object* x_60; -lean_inc(x_2); -x_56 = l_Lean_mkAppN(x_1, x_2); -x_57 = 0; -x_58 = 1; -x_59 = 1; -x_60 = l_Lean_Meta_mkLambdaFVars(x_2, x_56, x_57, x_58, x_59, x_6, x_7, x_8, x_9, x_46); +uint8_t x_47; +lean_dec(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); -return x_60; +lean_dec(x_3); +lean_dec(x_1); +x_47 = !lean_is_exclusive(x_43); +if (x_47 == 0) +{ +return x_43; } else { -lean_object* x_61; lean_object* x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; lean_object* x_66; -x_61 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5; -x_62 = l_Lean_Expr_app___override(x_1, x_61); -x_63 = 0; -x_64 = 1; -x_65 = 1; -x_66 = l_Lean_Meta_mkLambdaFVars(x_2, x_62, x_63, x_64, x_65, x_6, x_7, x_8, x_9, x_46); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -return x_66; +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_67; +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_array_fget(x_26, x_3); +lean_dec(x_26); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +x_52 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_52, 0, x_31); +lean_ctor_set(x_52, 1, x_32); +lean_ctor_set(x_52, 2, x_33); +lean_ctor_set(x_52, 3, x_51); +lean_ctor_set(x_52, 4, x_35); +lean_ctor_set_uint8(x_52, sizeof(void*)*5, x_34); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_53 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(x_39, x_1, x_30, x_52, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_53) == 0) +{ +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_53, 0); +lean_inc(x_54); +x_55 = lean_ctor_get(x_53, 1); +lean_inc(x_55); +lean_dec(x_53); +x_56 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_56, 0, x_54); +x_19 = x_56; +x_20 = x_55; +goto block_24; +} +else +{ +uint8_t x_57; +lean_dec(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_2); +lean_dec(x_3); lean_dec(x_1); -x_67 = !lean_is_exclusive(x_44); -if (x_67 == 0) +x_57 = !lean_is_exclusive(x_53); +if (x_57 == 0) { -return x_44; +return x_53; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_44, 0); -x_69 = lean_ctor_get(x_44, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_44); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; -} +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_53, 0); +x_59 = lean_ctor_get(x_53, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_53); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; } } } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; -x_11 = lean_unsigned_to_nat(0u); -x_12 = lean_nat_dec_eq(x_2, x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_sub(x_2, x_13); -lean_dec(x_2); -x_15 = 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_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_Expr_isLambda(x_16); -if (x_18 == 0) +else { -lean_object* x_19; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_16); -x_19 = lean_infer_type(x_16, x_6, x_7, x_8, x_9, x_17); -if (lean_obj_tag(x_19) == 0) +lean_object* x_61; +x_61 = lean_array_fget(x_27, x_3); +lean_dec(x_27); +if (x_37 == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); -lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___boxed), 10, 1); -lean_closure_set(x_22, 0, x_16); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_26); +x_62 = l_Lean_instInhabitedSubExpr; +x_63 = l___private_Init_GetElem_0__outOfBounds___rarg(x_62); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +x_64 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_64, 0, x_31); +lean_ctor_set(x_64, 1, x_32); +lean_ctor_set(x_64, 2, x_33); +lean_ctor_set(x_64, 3, x_63); +lean_ctor_set(x_64, 4, x_35); +lean_ctor_set_uint8(x_64, sizeof(void*)*5, x_34); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_23 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg(x_20, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_21); -if (lean_obj_tag(x_23) == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -x_24 = lean_ctor_get(x_4, 3); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_23, 1); -lean_inc(x_26); -lean_dec(x_23); -x_27 = !lean_is_exclusive(x_4); -if (x_27 == 0) -{ -lean_object* x_28; uint8_t x_29; -x_28 = lean_ctor_get(x_4, 3); -lean_dec(x_28); -x_29 = !lean_is_exclusive(x_24); -if (x_29 == 0) -{ -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_24, 0); -lean_dec(x_30); -lean_ctor_set(x_24, 0, x_25); -x_31 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -return x_31; -} -else -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_24, 1); -lean_inc(x_32); -lean_dec(x_24); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_25); -lean_ctor_set(x_33, 1, x_32); -lean_ctor_set(x_4, 3, x_33); -x_34 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_26); -return x_34; -} -} -else +lean_inc(x_1); +x_65 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(x_61, x_1, x_30, x_64, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_35 = lean_ctor_get(x_4, 0); -x_36 = lean_ctor_get(x_4, 1); -x_37 = lean_ctor_get(x_4, 2); -x_38 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); -x_39 = lean_ctor_get(x_4, 4); -lean_inc(x_39); -lean_inc(x_37); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_4); -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_24); - x_41 = lean_box(0); -} -if (lean_is_scalar(x_41)) { - x_42 = lean_alloc_ctor(0, 2, 0); -} else { - x_42 = x_41; -} -lean_ctor_set(x_42, 0, x_25); -lean_ctor_set(x_42, 1, x_40); -x_43 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_43, 0, x_35); -lean_ctor_set(x_43, 1, x_36); -lean_ctor_set(x_43, 2, x_37); -lean_ctor_set(x_43, 3, x_42); -lean_ctor_set(x_43, 4, x_39); -lean_ctor_set_uint8(x_43, sizeof(void*)*5, x_38); -x_44 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_14, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_26); -return x_44; -} +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_68, 0, x_66); +x_19 = x_68; +x_20 = x_67; +goto block_24; } else { -uint8_t x_45; -lean_dec(x_14); +uint8_t x_69; +lean_dec(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); lean_dec(x_3); lean_dec(x_1); -x_45 = !lean_is_exclusive(x_23); -if (x_45 == 0) +x_69 = !lean_is_exclusive(x_65); +if (x_69 == 0) { -return x_23; +return x_65; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_23, 0); -x_47 = lean_ctor_get(x_23, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_23); -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_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_65, 0); +x_71 = lean_ctor_get(x_65, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_65); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; } } } else { -uint8_t x_49; -lean_dec(x_16); -lean_dec(x_14); +lean_object* x_73; lean_object* x_74; lean_object* x_75; +x_73 = lean_array_fget(x_26, x_3); +lean_dec(x_26); +lean_inc(x_35); +lean_inc(x_33); +lean_inc(x_32); +lean_inc(x_31); +x_74 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_74, 0, x_31); +lean_ctor_set(x_74, 1, x_32); +lean_ctor_set(x_74, 2, x_33); +lean_ctor_set(x_74, 3, x_73); +lean_ctor_set(x_74, 4, x_35); +lean_ctor_set_uint8(x_74, sizeof(void*)*5, x_34); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_1); +x_75 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(x_61, x_1, x_30, x_74, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_78 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_78, 0, x_76); +x_19 = x_78; +x_20 = x_77; +goto block_24; +} +else +{ +uint8_t x_79; +lean_dec(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); lean_dec(x_3); lean_dec(x_1); -x_49 = !lean_is_exclusive(x_19); -if (x_49 == 0) +x_79 = !lean_is_exclusive(x_75); +if (x_79 == 0) { -return x_19; +return x_75; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_19, 0); -x_51 = lean_ctor_get(x_19, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_19); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_75, 0); +x_81 = lean_ctor_get(x_75, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_75); +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; } } } -else +} +block_24: { -lean_object* x_53; -lean_dec(x_16); -x_53 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -return x_53; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +x_22 = lean_nat_add(x_3, x_5); +lean_dec(x_3); +x_2 = x_18; +x_3 = x_22; +x_6 = x_21; +x_13 = x_20; +goto _start; } } else { -lean_object* x_54; +lean_object* x_83; +lean_dec(x_12); +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_54 = lean_apply_8(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_54; -} +lean_dec(x_1); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_6); +lean_ctor_set(x_83, 1, x_13); +return x_83; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 10, 0); -return x_2; +lean_object* x_84; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_84 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_84, 0, x_6); +lean_ctor_set(x_84, 1, x_13); +return x_84; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_object* x_10) { +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__16(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { uint8_t x_11; @@ -22664,49 +23029,87 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__15(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3___rarg), 10, 0); -return x_2; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appFn_x21(x_10); +x_13 = lean_unsigned_to_nat(0u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__16(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___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_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_12; +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_inc(x_9); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = l_Lean_Expr_appArg_x21(x_10); +lean_dec(x_10); +x_13 = lean_unsigned_to_nat(1u); +x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__16(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); +lean_dec(x_10); +x_13 = l_Lean_Expr_isApp(x_11); +lean_dec(x_11); +if (x_13 == 0) +{ +lean_object* x_14; +lean_dec(x_2); +x_14 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +return x_14; +} +else +{ +lean_object* x_15; lean_object* x_16; +lean_inc(x_2); +x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14), 9, 2); +lean_closure_set(x_15, 0, x_1); +lean_closure_set(x_15, 1, x_2); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_12 = lean_apply_8(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_12) == 0) +lean_inc(x_3); +x_16 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__15(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_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_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_Expr_bindingBody_x21(x_2); -lean_dec(x_2); -x_16 = lean_expr_instantiate1(x_15, x_4); -lean_dec(x_4); -lean_dec(x_15); -x_17 = lean_apply_1(x_3, x_13); -x_18 = lean_unsigned_to_nat(1u); -x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__3___rarg(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -return x_19; +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_apply_1(x_2, x_17); +x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__17(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_20; } else { -uint8_t x_20; -lean_dec(x_10); -lean_dec(x_9); +uint8_t x_21; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -22714,593 +23117,439 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_20 = !lean_is_exclusive(x_12); -if (x_20 == 0) +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) { -return x_12; +return x_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_12, 0); -x_22 = lean_ctor_get(x_12, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_12); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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) { -_start: -{ -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; -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_14 = l_Lean_Expr_binderInfo(x_12); -x_15 = l_Lean_Expr_bindingDomain_x21(x_12); -x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg___lambda__1), 11, 3); -lean_closure_set(x_16, 0, x_2); -lean_closure_set(x_16, 1, x_12); -lean_closure_set(x_16, 2, x_3); -x_17 = 0; -x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___spec__3___rarg(x_1, x_14, x_15, x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg), 10, 0); -return x_3; -} -} -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(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__2___boxed), 9, 1); -lean_closure_set(x_10, 0, x_2); -x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1; -x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; -} -} -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(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg), 9, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___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_object* x_11) { -_start: -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = l_Lean_Syntax_getId(x_4); -lean_dec(x_4); -x_13 = lean_array_push(x_1, x_12); -x_14 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(x_2, x_3, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -return x_14; -} -} -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -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_14 = l_Lean_Expr_bindingName_x21(x_12); -lean_dec(x_12); -x_15 = lean_erase_macro_scopes(x_14); -x_16 = l_Array_contains___at_Lean_findField_x3f___spec__1(x_3, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -lean_dec(x_15); -x_17 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg___lambda__1), 11, 3); -lean_closure_set(x_17, 0, x_3); -lean_closure_set(x_17, 1, x_1); -lean_closure_set(x_17, 2, x_2); -x_18 = l_Lean_PrettyPrinter_Delaborator_withBindingBodyUnusedName___rarg(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_18; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_6, 1); -lean_inc(x_19); -x_20 = l_Lean_LocalContext_getUnusedName(x_19, x_15); -lean_inc(x_20); -x_21 = lean_array_push(x_3, x_20); -x_22 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg), 10, 3); -lean_closure_set(x_22, 0, x_1); -lean_closure_set(x_22, 1, x_2); -lean_closure_set(x_22, 2, x_21); -x_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg(x_20, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -return x_23; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___rarg), 10, 0); -return x_2; -} +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___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) { -_start: -{ -lean_object* x_11; -x_11 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; lean_object* x_11; -x_10 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_11 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg(x_2, x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg), 9, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) { -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); +uint8_t x_10; x_10 = !lean_is_exclusive(x_9); if (x_10 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; +lean_object* x_11; lean_object* x_12; x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_12, x_1); -lean_ctor_set(x_9, 0, x_13); +x_12 = lean_array_push(x_1, x_11); +lean_ctor_set(x_9, 0, x_12); return x_9; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = lean_ctor_get(x_9, 0); -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); lean_inc(x_14); +lean_inc(x_13); lean_dec(x_9); -x_16 = lean_ctor_get(x_14, 0); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Meta_Match_Extension_getMatcherInfo_x3f(x_16, x_1); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_17); -lean_ctor_set(x_18, 1, x_15); -return x_18; -} +x_15 = lean_array_push(x_1, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: +else { -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) +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 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; +return x_9; } 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_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } -static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unknown constant '", 18); -return x_1; } } -static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2() { +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); +x_1 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_2 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_PrettyPrinter_Delaborator_instMonadQuotationDelabM___spec__1___rarg___boxed), 8, 1); +lean_closure_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3() { +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("'", 1); +x_1 = lean_alloc_closure((void*)(l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___lambda__1), 8, 0); return x_1; } } -static lean_object* _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { _start: { -lean_object* x_9; 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_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_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_5, x_18); +lean_dec(x_5); +x_20 = lean_array_fget(x_4, x_6); +lean_inc(x_14); lean_inc(x_13); -lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_9); -x_15 = lean_box(0); -x_16 = l_Lean_Expr_const___override(x_1, x_15); -x_17 = l_Lean_MessageData_ofExpr(x_16); -x_18 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__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_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__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_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -return x_22; -} -else +lean_inc(x_12); +lean_inc(x_11); +x_21 = lean_infer_type(x_20, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; -lean_dec(x_1); -x_23 = lean_ctor_get(x_14, 0); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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; 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_14); -lean_ctor_set(x_9, 0, x_23); -return x_9; -} -} -else +lean_dec(x_21); +x_24 = lean_ctor_get(x_1, 6); +x_25 = lean_array_get_size(x_24); +x_26 = lean_nat_dec_lt(x_6, x_25); +lean_dec(x_25); +x_27 = l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1; +x_28 = l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2; +x_29 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14), 9, 2); +lean_closure_set(x_29, 0, x_27); +lean_closure_set(x_29, 1, x_28); +x_30 = lean_ctor_get(x_9, 0); +lean_inc(x_30); +x_31 = lean_ctor_get(x_9, 1); +lean_inc(x_31); +x_32 = lean_ctor_get(x_9, 2); +lean_inc(x_32); +x_33 = lean_ctor_get_uint8(x_9, sizeof(void*)*5); +x_34 = lean_ctor_get(x_9, 3); +lean_inc(x_34); +x_35 = lean_ctor_get(x_9, 4); +lean_inc(x_35); +x_36 = lean_nat_add(x_2, x_6); +x_37 = lean_ctor_get(x_34, 1); +lean_inc(x_37); +lean_dec(x_34); +x_38 = l_Lean_SubExpr_Pos_pushNthBindingDomain(x_36, x_37); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_22); +lean_ctor_set(x_39, 1, x_38); +x_40 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_40, 0, x_30); +lean_ctor_set(x_40, 1, x_31); +lean_ctor_set(x_40, 2, x_32); +lean_ctor_set(x_40, 3, x_39); +lean_ctor_set(x_40, 4, x_35); +lean_ctor_set_uint8(x_40, sizeof(void*)*5, x_33); +if (x_26 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -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_ctor_get(x_24, 0); -lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); -if (lean_obj_tag(x_27) == 0) +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_42 = l___private_Init_GetElem_0__outOfBounds___rarg(x_41); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_43 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___rarg(x_42, x_29, x_16, x_40, x_10, x_11, x_12, x_13, x_14, x_23); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_28 = lean_box(0); -x_29 = l_Lean_Expr_const___override(x_1, x_28); -x_30 = l_Lean_MessageData_ofExpr(x_29); -x_31 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__3(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_25); -return x_35; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +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_nat_add(x_6, x_18); +lean_dec(x_6); +x_47 = lean_array_push(x_8, x_44); +x_5 = x_19; +x_6 = x_46; +x_7 = lean_box(0); +x_8 = x_47; +x_15 = x_45; +goto _start; } else { -lean_object* x_36; lean_object* x_37; -lean_dec(x_1); -x_36 = lean_ctor_get(x_27, 0); -lean_inc(x_36); -lean_dec(x_27); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_25); -return x_37; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___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, lean_object* x_10) { -_start: -{ -lean_object* x_11; uint8_t x_12; lean_object* x_13; -x_11 = lean_alloc_closure((void*)(l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__6___rarg___lambda__1), 10, 3); -lean_closure_set(x_11, 0, x_2); -lean_closure_set(x_11, 1, x_4); -lean_closure_set(x_11, 2, x_5); -x_12 = 0; -x_13 = l___private_Lean_Meta_Basic_0__Lean_Meta_lambdaTelescopeImp___rarg(x_1, x_12, x_11, x_3, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_13) == 0) -{ -uint8_t x_14; -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) +uint8_t x_49; +lean_dec(x_19); +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_6); +x_49 = !lean_is_exclusive(x_43); +if (x_49 == 0) { -return x_13; +return x_43; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 0); -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_15); -lean_ctor_set(x_17, 1, x_16); -return x_17; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_43, 0); +x_51 = lean_ctor_get(x_43, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_43); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; +} } } else { -uint8_t x_18; -x_18 = !lean_is_exclusive(x_13); -if (x_18 == 0) +lean_object* x_53; lean_object* x_54; +x_53 = lean_array_fget(x_24, x_6); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_11); +lean_inc(x_10); +x_54 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch_usingNames___rarg(x_53, x_29, x_16, x_40, x_10, x_11, x_12, x_13, x_14, x_23); +if (lean_obj_tag(x_54) == 0) { -return x_13; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_nat_add(x_6, x_18); +lean_dec(x_6); +x_58 = lean_array_push(x_8, x_55); +x_5 = x_19; +x_6 = x_57; +x_7 = lean_box(0); +x_8 = x_58; +x_15 = x_56; +goto _start; } 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); +uint8_t x_60; +lean_dec(x_19); +lean_dec(x_14); 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_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +x_60 = !lean_is_exclusive(x_54); +if (x_60 == 0) +{ +return x_54; } +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_54, 0); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_54); +x_63 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_63, 0, x_61); +lean_ctor_set(x_63, 1, x_62); +return x_63; } } -LEAN_EXPORT lean_object* l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___rarg___boxed), 10, 0); -return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: -{ -uint8_t x_11; -x_11 = !lean_is_exclusive(x_4); -if (x_11 == 0) +else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_12 = lean_ctor_get(x_4, 3); -x_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_SubExpr_Pos_push(x_13, x_2); +uint8_t x_64; +lean_dec(x_19); +lean_dec(x_14); lean_dec(x_13); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_14); -lean_ctor_set(x_4, 3, x_15); -x_16 = lean_apply_7(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_16; +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +x_64 = !lean_is_exclusive(x_21); +if (x_64 == 0) +{ +return x_21; } 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; -x_17 = lean_ctor_get(x_4, 0); -x_18 = lean_ctor_get(x_4, 1); -x_19 = lean_ctor_get(x_4, 2); -x_20 = lean_ctor_get_uint8(x_4, sizeof(void*)*5); -x_21 = lean_ctor_get(x_4, 3); -x_22 = lean_ctor_get(x_4, 4); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_4); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); +lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_65 = lean_ctor_get(x_21, 0); +x_66 = lean_ctor_get(x_21, 1); +lean_inc(x_66); +lean_inc(x_65); lean_dec(x_21); -x_24 = l_Lean_SubExpr_Pos_push(x_23, x_2); -lean_dec(x_23); -x_25 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_25, 0, x_1); -lean_ctor_set(x_25, 1, x_24); -x_26 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_26, 0, x_17); -lean_ctor_set(x_26, 1, x_18); -lean_ctor_set(x_26, 2, x_19); -lean_ctor_set(x_26, 3, x_25); -lean_ctor_set(x_26, 4, x_22); -lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_20); -x_27 = lean_apply_7(x_3, x_26, x_5, x_6, x_7, x_8, x_9, x_10); -return x_27; +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_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: +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; -x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); +lean_object* x_68; +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_12 = l_Lean_Expr_appFn_x21(x_10); -x_13 = lean_unsigned_to_nat(0u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__7(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_14; +lean_dec(x_6); +lean_dec(x_5); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_8); +lean_ctor_set(x_68, 1, x_15); +return x_68; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1() { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Expr_appArg_x21(x_10); -lean_dec(x_10); -x_13 = lean_unsigned_to_nat(1u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__7(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_14; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("matchDiscr", 10); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2() { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_3, x_4, x_5, x_6, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_Expr_isApp(x_11); -lean_dec(x_11); -if (x_13 == 0) +lean_object* x_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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1; +x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19(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: { -lean_object* x_14; -lean_dec(x_2); -x_14 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -return x_14; +uint8_t x_11; +x_11 = lean_usize_dec_lt(x_2, x_1); +if (x_11 == 0) +{ +lean_object* x_12; +lean_dec(x_8); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_3); +lean_ctor_set(x_12, 1, x_10); +return x_12; } else { -lean_object* x_15; lean_object* x_16; -lean_inc(x_2); -x_15 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5), 9, 2); -lean_closure_set(x_15, 0, x_1); -lean_closure_set(x_15, 1, x_2); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_16 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__6(x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_array_uget(x_3, x_2); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_3, x_2, x_14); +x_16 = lean_ctor_get(x_13, 0); +lean_inc(x_16); if (lean_obj_tag(x_16) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_16, 0); +lean_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; size_t x_26; size_t x_27; lean_object* x_28; +x_17 = lean_ctor_get(x_13, 1); lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); +lean_dec(x_13); +x_18 = lean_ctor_get(x_8, 5); lean_inc(x_18); -lean_dec(x_16); -x_19 = lean_apply_1(x_2, x_17); -x_20 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__8(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -return x_20; -} -else -{ -uint8_t x_21; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) -{ -return x_16; +x_19 = 0; +x_20 = l_Lean_SourceInfo_fromRef(x_18, x_19); +x_21 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_22 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_inc(x_20); +x_23 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_23, 0, x_20); +lean_ctor_set(x_23, 1, x_21); +lean_ctor_set(x_23, 2, x_22); +x_24 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2; +x_25 = l_Lean_Syntax_node2(x_20, x_24, x_23, x_17); +x_26 = 1; +x_27 = lean_usize_add(x_2, x_26); +x_28 = lean_array_uset(x_15, x_2, x_25); +x_2 = x_27; +x_3 = x_28; +goto _start; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_inc(x_22); +lean_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; lean_object* x_40; lean_object* x_41; size_t x_42; size_t x_43; lean_object* x_44; +x_30 = lean_ctor_get(x_13, 1); +lean_inc(x_30); +lean_dec(x_13); +x_31 = lean_ctor_get(x_16, 0); +lean_inc(x_31); lean_dec(x_16); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} +x_32 = lean_ctor_get(x_8, 5); +lean_inc(x_32); +x_33 = 0; +x_34 = l_Lean_SourceInfo_fromRef(x_32, x_33); +x_35 = lean_mk_syntax_ident(x_31); +x_36 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; +lean_inc(x_34); +x_37 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_36); +x_38 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_34); +x_39 = l_Lean_Syntax_node2(x_34, x_38, x_35, x_37); +x_40 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2; +x_41 = l_Lean_Syntax_node2(x_34, x_40, x_39, x_30); +x_42 = 1; +x_43 = lean_usize_add(x_2, x_42); +x_44 = lean_array_uset(x_15, x_2, x_41); +x_2 = x_43; +x_3 = x_44; +goto _start; } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(size_t x_1, size_t x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(size_t x_1, size_t x_2, lean_object* x_3) { _start: { uint8_t x_4; @@ -23324,7 +23573,7 @@ goto _start; } } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__1() { _start: { lean_object* x_1; @@ -23332,19 +23581,19 @@ x_1 = lean_mk_string_from_bytes("matchAlt", 8); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___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_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1; +x_4 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___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_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3() { _start: { lean_object* x_1; @@ -23352,7 +23601,7 @@ x_1 = lean_mk_string_from_bytes("|", 1); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4() { +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4() { _start: { lean_object* x_1; @@ -23360,7 +23609,7 @@ x_1 = lean_mk_string_from_bytes("=>", 2); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21(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; @@ -23385,7 +23634,7 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_9, 1); lean_inc(x_15); lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; +x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3; lean_inc(x_2); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_2); @@ -23404,12 +23653,12 @@ lean_ctor_set(x_21, 2, x_20); lean_inc(x_3); lean_inc(x_2); x_22 = l_Lean_Syntax_node1(x_2, x_3, x_21); -x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; +x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; lean_inc(x_2); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_2); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; +x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2; lean_inc(x_2); x_26 = l_Lean_Syntax_node4(x_2, x_25, x_17, x_22, x_24, x_15); x_27 = lean_array_uset(x_11, x_6, x_26); @@ -23419,7 +23668,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22(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; @@ -23444,7 +23693,7 @@ lean_inc(x_14); x_15 = lean_ctor_get(x_9, 1); lean_inc(x_15); lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; +x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3; lean_inc(x_2); x_17 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_17, 0, x_2); @@ -23463,12 +23712,12 @@ lean_ctor_set(x_21, 2, x_20); lean_inc(x_3); lean_inc(x_2); x_22 = l_Lean_Syntax_node1(x_2, x_3, x_21); -x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; +x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; lean_inc(x_2); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_2); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; +x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2; lean_inc(x_2); x_26 = l_Lean_Syntax_node4(x_2, x_25, x_17, x_22, x_24, x_15); x_27 = lean_array_uset(x_11, x_6, x_26); @@ -23478,7 +23727,7 @@ goto _start; } } } -static lean_object* _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1() { +static lean_object* _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23490,341 +23739,305 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(lean_object* x_1) { +LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1; +x_2 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1; x_3 = lean_panic_fn(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_12 = lean_unsigned_to_nat(0u); +x_13 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_4, x_12); +x_14 = l_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___lambda__3___closed__1; +lean_inc(x_13); +x_15 = lean_mk_array(x_13, x_14); +x_16 = lean_unsigned_to_nat(1u); +x_17 = lean_nat_sub(x_13, x_16); +lean_dec(x_13); +x_18 = l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(x_4, x_15, x_17); +x_19 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_19) == 0) { -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; -} -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; +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); +x_22 = l_Lean_Core_instantiateTypeLevelParams(x_20, x_2, x_9, x_10, x_21); +lean_dec(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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_25 = l___private_Lean_Meta_Basic_0__Lean_Meta_instantiateForallAux(x_18, x_12, x_23, x_7, x_8, x_9, x_10, x_24); +lean_dec(x_18); +if (lean_obj_tag(x_25) == 0) { -lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_9 = lean_array_uget(x_7, x_6); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_7, x_6, x_10); -x_12 = 1; -x_13 = lean_usize_add(x_6, x_12); -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; -lean_inc(x_2); -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_19 = l_Lean_Syntax_SepArray_ofElems(x_18, x_14); -lean_dec(x_14); -lean_inc(x_4); -x_20 = l_Array_append___rarg(x_4, x_19); -lean_inc(x_3); -lean_inc(x_2); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_2); -lean_ctor_set(x_21, 1, x_3); -lean_ctor_set(x_21, 2, x_20); -lean_inc(x_3); -lean_inc(x_2); -x_22 = l_Lean_Syntax_node1(x_2, x_3, x_21); -x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; -lean_inc(x_2); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_2); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; -lean_inc(x_2); -x_26 = l_Lean_Syntax_node4(x_2, x_25, x_17, x_22, x_24, x_15); -x_27 = lean_array_uset(x_11, x_6, x_26); -x_6 = x_13; -x_7 = x_27; -goto _start; -} -} -} -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, size_t x_5, size_t x_6, lean_object* x_7) { -_start: +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_26 = lean_ctor_get(x_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_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_getOptionsAtCurrPos___spec__1(x_5, x_6, x_7, x_8, x_9, x_10, x_27); +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = l_Lean_SubExpr_Pos_typeCoord; +x_32 = l_Lean_SubExpr_Pos_push(x_29, x_31); +lean_dec(x_29); +lean_inc(x_26); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_26); +lean_ctor_set(x_33, 1, x_32); +x_34 = 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_30); +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); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +x_37 = lean_infer_type(x_35, x_7, x_8, x_9, x_10, x_36); +if (lean_obj_tag(x_37) == 0) { -uint8_t x_8; -x_8 = lean_usize_dec_lt(x_6, x_5); -if (x_8 == 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_40 = l_Lean_Meta_isExprDefEq(x_26, x_38, x_7, x_8, x_9, x_10, x_39); +if (lean_obj_tag(x_40) == 0) { -lean_dec(x_4); +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; lean_object* x_44; uint8_t x_45; +lean_dec(x_33); lean_dec(x_3); -lean_dec(x_2); -return x_7; +x_43 = lean_ctor_get(x_40, 1); +lean_inc(x_43); +lean_dec(x_40); +x_44 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_43); +x_45 = !lean_is_exclusive(x_44); +if (x_45 == 0) +{ +return x_44; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_9 = lean_array_uget(x_7, x_6); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_array_uset(x_7, x_6, x_10); -x_12 = 1; -x_13 = lean_usize_add(x_6, x_12); -x_14 = lean_ctor_get(x_9, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_9, 1); -lean_inc(x_15); -lean_dec(x_9); -x_16 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3; -lean_inc(x_2); -x_17 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_17, 0, x_2); -lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; -x_19 = l_Lean_Syntax_SepArray_ofElems(x_18, x_14); -lean_dec(x_14); -lean_inc(x_4); -x_20 = l_Array_append___rarg(x_4, x_19); -lean_inc(x_3); -lean_inc(x_2); -x_21 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_21, 0, x_2); -lean_ctor_set(x_21, 1, x_3); -lean_ctor_set(x_21, 2, x_20); -lean_inc(x_3); -lean_inc(x_2); -x_22 = l_Lean_Syntax_node1(x_2, x_3, x_21); -x_23 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; -lean_inc(x_2); -x_24 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_24, 0, x_2); -lean_ctor_set(x_24, 1, x_23); -x_25 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2; -lean_inc(x_2); -x_26 = l_Lean_Syntax_node4(x_2, x_25, x_17, x_22, x_24, x_15); -x_27 = lean_array_uset(x_11, x_6, x_26); -x_6 = x_13; -x_7 = x_27; -goto _start; -} -} +lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_44, 0); +x_47 = lean_ctor_get(x_44, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_44); +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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -if (lean_obj_tag(x_1) == 4) -{ -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); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -lean_inc(x_9); -x_11 = l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_13; lean_object* x_14; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_7); -lean_dec(x_6); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_13); -return x_14; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_11, 1); -lean_inc(x_15); -lean_dec(x_11); -x_16 = lean_ctor_get(x_12, 0); -lean_inc(x_16); -lean_dec(x_12); -x_17 = l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_15); -if (lean_obj_tag(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_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Core_instantiateTypeLevelParams(x_18, x_10, x_6, x_7, x_19); -lean_dec(x_18); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_box(0); -x_24 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_25 = 0; -x_26 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_26, 0, x_16); -lean_ctor_set(x_26, 1, x_22); -lean_ctor_set(x_26, 2, x_24); -lean_ctor_set(x_26, 3, x_23); -lean_ctor_set(x_26, 4, x_24); -lean_ctor_set(x_26, 5, x_24); -lean_ctor_set(x_26, 6, x_24); -lean_ctor_set(x_26, 7, x_24); -lean_ctor_set_uint8(x_26, sizeof(void*)*8, x_25); -lean_ctor_set(x_20, 0, x_26); -return x_20; -} -else +uint8_t x_49; +x_49 = !lean_is_exclusive(x_40); +if (x_49 == 0) { -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; -x_27 = lean_ctor_get(x_20, 0); -x_28 = lean_ctor_get(x_20, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_20); -x_29 = lean_box(0); -x_30 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -x_31 = 0; -x_32 = lean_alloc_ctor(0, 8, 1); -lean_ctor_set(x_32, 0, x_16); -lean_ctor_set(x_32, 1, x_27); -lean_ctor_set(x_32, 2, x_30); -lean_ctor_set(x_32, 3, x_29); -lean_ctor_set(x_32, 4, x_30); -lean_ctor_set(x_32, 5, x_30); -lean_ctor_set(x_32, 6, x_30); -lean_ctor_set(x_32, 7, x_30); -lean_ctor_set_uint8(x_32, sizeof(void*)*8, 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_28); -return x_33; +lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; +x_50 = lean_ctor_get(x_40, 0); +lean_dec(x_50); +x_51 = lean_box(0); +x_52 = 0; +x_53 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_54 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_54, 0, x_3); +lean_ctor_set(x_54, 1, x_33); +lean_ctor_set(x_54, 2, x_51); +lean_ctor_set(x_54, 3, x_53); +lean_ctor_set(x_54, 4, x_53); +lean_ctor_set(x_54, 5, x_53); +lean_ctor_set(x_54, 6, x_53); +lean_ctor_set(x_54, 7, x_53); +lean_ctor_set_uint8(x_54, sizeof(void*)*8, x_52); +lean_ctor_set(x_40, 0, x_54); +return x_40; +} +else +{ +lean_object* x_55; lean_object* x_56; uint8_t x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_55 = lean_ctor_get(x_40, 1); +lean_inc(x_55); +lean_dec(x_40); +x_56 = lean_box(0); +x_57 = 0; +x_58 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +x_59 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_59, 0, x_3); +lean_ctor_set(x_59, 1, x_33); +lean_ctor_set(x_59, 2, x_56); +lean_ctor_set(x_59, 3, x_58); +lean_ctor_set(x_59, 4, x_58); +lean_ctor_set(x_59, 5, x_58); +lean_ctor_set(x_59, 6, x_58); +lean_ctor_set(x_59, 7, x_58); +lean_ctor_set_uint8(x_59, sizeof(void*)*8, x_57); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_55); +return x_60; +} } } else { -uint8_t x_34; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_7); -lean_dec(x_6); -x_34 = !lean_is_exclusive(x_17); -if (x_34 == 0) +uint8_t x_61; +lean_dec(x_33); +lean_dec(x_3); +x_61 = !lean_is_exclusive(x_40); +if (x_61 == 0) { -return x_17; +return x_40; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_17, 0); -x_36 = lean_ctor_get(x_17, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_17); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; -} +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_40, 0); +x_63 = lean_ctor_get(x_40, 1); +lean_inc(x_63); +lean_inc(x_62); +lean_dec(x_40); +x_64 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_64, 0, x_62); +lean_ctor_set(x_64, 1, x_63); +return x_64; } } } else { -lean_object* x_38; +uint8_t x_65; +lean_dec(x_33); +lean_dec(x_26); +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_38 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -return x_38; -} -} +lean_dec(x_3); +x_65 = !lean_is_exclusive(x_37); +if (x_65 == 0) +{ +return x_37; } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: +else { -uint8_t x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; -x_10 = 0; -x_11 = 1; -x_12 = 1; -x_13 = l_Lean_Meta_mkForallFVars(x_1, x_2, x_10, x_11, x_12, x_5, x_6, x_7, x_8, x_9); -return x_13; +lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_66 = lean_ctor_get(x_37, 0); +x_67 = lean_ctor_get(x_37, 1); +lean_inc(x_67); +lean_inc(x_66); +lean_dec(x_37); +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_PrettyPrinter_Delaborator_delabAppMatch___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; -x_9 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +} +else { -uint8_t x_10; -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) +uint8_t x_69; +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +x_69 = !lean_is_exclusive(x_25); +if (x_69 == 0) { -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_1); -lean_ctor_set(x_12, 1, x_11); -lean_ctor_set(x_9, 0, x_12); -return x_9; +return x_25; } else { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_13 = lean_ctor_get(x_9, 0); -x_14 = lean_ctor_get(x_9, 1); -lean_inc(x_14); -lean_inc(x_13); -lean_dec(x_9); -x_15 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_15, 0, x_1); -lean_ctor_set(x_15, 1, x_13); -x_16 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -return x_16; +lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_70 = lean_ctor_get(x_25, 0); +x_71 = lean_ctor_get(x_25, 1); +lean_inc(x_71); +lean_inc(x_70); +lean_dec(x_25); +x_72 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_72, 0, x_70); +lean_ctor_set(x_72, 1, x_71); +return x_72; +} } } else { -uint8_t x_17; -lean_dec(x_1); -x_17 = !lean_is_exclusive(x_9); -if (x_17 == 0) +uint8_t x_73; +lean_dec(x_18); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_3); +lean_dec(x_2); +x_73 = !lean_is_exclusive(x_19); +if (x_73 == 0) { -return x_9; +return x_19; } else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_9, 0); -x_19 = lean_ctor_get(x_9, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_9); -x_20 = lean_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_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_19, 0); +x_75 = lean_ctor_get(x_19, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_19); +x_76 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_76, 0, x_74); +lean_ctor_set(x_76, 1, x_75); +return x_76; } } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: +{ +uint8_t x_10; uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_10 = 0; +x_11 = 1; +x_12 = 1; +x_13 = l_Lean_Meta_mkForallFVars(x_1, x_2, x_10, x_11, x_12, x_5, x_6, x_7, x_8, x_9); +return x_13; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -23832,977 +24045,603 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch_ return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3), 8, 0); +x_1 = lean_mk_string_from_bytes("Lean.PrettyPrinter.Delaborator.delabAppMatch", 44); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("matchDiscr", 10); +x_1 = lean_mk_string_from_bytes("impossible, number of arguments does not match arity", 52); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_1 = l_Lean_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_delabAppMatch___lambda__4___closed__3; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2; +x_3 = lean_unsigned_to_nat(674u); +x_4 = lean_unsigned_to_nat(10u); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_9 = lean_ctor_get(x_1, 0); +lean_object* x_9; +x_9 = lean_ctor_get(x_1, 2); lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -x_11 = lean_ctor_get(x_1, 2); +if (lean_obj_tag(x_9) == 0) +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +x_10 = 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_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); -x_12 = lean_ctor_get(x_1, 3); +x_12 = lean_ctor_get(x_10, 1); lean_inc(x_12); -x_13 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -x_14 = lean_ctor_get(x_1, 4); -lean_inc(x_14); -x_15 = lean_ctor_get(x_1, 5); -lean_inc(x_15); -x_16 = lean_ctor_get(x_1, 6); +lean_dec(x_10); +x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1; +x_14 = 0; +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_11); +x_15 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___rarg(x_11, x_13, x_14, x_2, x_3, x_4, x_5, x_6, x_7, 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; lean_object* x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -x_17 = lean_ctor_get(x_1, 7); +x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); -if (lean_is_exclusive(x_1)) { - lean_ctor_release(x_1, 0); - lean_ctor_release(x_1, 1); - lean_ctor_release(x_1, 2); - lean_ctor_release(x_1, 3); - lean_ctor_release(x_1, 4); - lean_ctor_release(x_1, 5); - lean_ctor_release(x_1, 6); - lean_ctor_release(x_1, 7); - x_18 = x_1; -} else { - lean_dec_ref(x_1); - x_18 = lean_box(0); -} -x_19 = lean_array_get_size(x_11); -x_20 = lean_ctor_get(x_9, 0); +lean_dec(x_15); +x_18 = lean_ctor_get(x_2, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_2, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_2, 2); lean_inc(x_20); -x_21 = lean_nat_dec_lt(x_19, x_20); -lean_dec(x_20); -lean_dec(x_19); -if (x_21 == 0) -{ -if (lean_obj_tag(x_12) == 0) -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_22 = 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_23 = lean_ctor_get(x_22, 0); +x_21 = lean_ctor_get_uint8(x_2, sizeof(void*)*5); +x_22 = lean_ctor_get(x_2, 3); +lean_inc(x_22); +x_23 = lean_ctor_get(x_22, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); -lean_inc(x_24); lean_dec(x_22); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1; -x_26 = 0; +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +x_25 = lean_ctor_get(x_2, 4); +lean_inc(x_25); +x_26 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_26, 0, x_18); +lean_ctor_set(x_26, 1, x_19); +lean_ctor_set(x_26, 2, x_20); +lean_ctor_set(x_26, 3, x_24); +lean_ctor_set(x_26, 4, x_25); +lean_ctor_set_uint8(x_26, sizeof(void*)*5, x_21); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_23); -x_27 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__4___rarg(x_23, x_25, x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_24); +x_27 = l_Lean_PrettyPrinter_Delaborator_delab(x_26, x_3, x_4, x_5, x_6, x_7, x_17); if (lean_obj_tag(x_27) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; lean_object* x_35; uint8_t x_36; -x_28 = lean_ctor_get(x_2, 3); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 0); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); lean_dec(x_27); -x_31 = lean_ctor_get(x_2, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_2, 1); -lean_inc(x_32); -x_33 = lean_ctor_get(x_2, 2); -lean_inc(x_33); -x_34 = lean_ctor_get_uint8(x_2, sizeof(void*)*5); -x_35 = lean_ctor_get(x_2, 4); -lean_inc(x_35); -x_36 = !lean_is_exclusive(x_28); -if (x_36 == 0) +x_30 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2; +x_31 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_29); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_28, 0); -lean_dec(x_37); -lean_ctor_set(x_28, 0, x_29); -x_38 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_38, 0, x_31); -lean_ctor_set(x_38, 1, x_32); -lean_ctor_set(x_38, 2, x_33); -lean_ctor_set(x_38, 3, x_28); -lean_ctor_set(x_38, 4, x_35); -lean_ctor_set_uint8(x_38, sizeof(void*)*5, x_34); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_39 = l_Lean_PrettyPrinter_Delaborator_delab(x_38, x_3, x_4, x_5, x_6, x_7, x_30); -if (lean_obj_tag(x_39) == 0) +uint8_t x_32; +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 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_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; +x_33 = lean_ctor_get(x_31, 0); +x_34 = lean_ctor_get(x_1, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_1, 1); +lean_inc(x_35); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_28); +lean_ctor_set(x_36, 1, x_11); +x_37 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_37, 0, x_36); +x_38 = lean_ctor_get(x_1, 3); +lean_inc(x_38); +x_39 = lean_ctor_get(x_1, 4); +lean_inc(x_39); +x_40 = lean_ctor_get(x_1, 5); lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +x_41 = lean_ctor_get(x_1, 6); lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4; -x_43 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_41); -if (lean_obj_tag(x_43) == 0) -{ -uint8_t x_44; -x_44 = !lean_is_exclusive(x_43); -if (x_44 == 0) -{ -lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; -x_45 = lean_ctor_get(x_43, 0); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_40); -lean_ctor_set(x_46, 1, x_23); -x_47 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_47, 0, x_46); -if (lean_is_scalar(x_18)) { - x_48 = lean_alloc_ctor(0, 8, 1); -} else { - x_48 = x_18; -} -lean_ctor_set(x_48, 0, x_9); -lean_ctor_set(x_48, 1, x_10); -lean_ctor_set(x_48, 2, x_11); -lean_ctor_set(x_48, 3, x_47); -lean_ctor_set(x_48, 4, x_14); -lean_ctor_set(x_48, 5, x_15); -lean_ctor_set(x_48, 6, x_16); -lean_ctor_set(x_48, 7, x_17); -x_49 = lean_unbox(x_45); -lean_dec(x_45); -lean_ctor_set_uint8(x_48, sizeof(void*)*8, x_49); -lean_ctor_set(x_43, 0, x_48); -return x_43; +x_42 = lean_ctor_get(x_1, 7); +lean_inc(x_42); +lean_dec(x_1); +x_43 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_43, 0, x_34); +lean_ctor_set(x_43, 1, x_35); +lean_ctor_set(x_43, 2, x_37); +lean_ctor_set(x_43, 3, x_38); +lean_ctor_set(x_43, 4, x_39); +lean_ctor_set(x_43, 5, x_40); +lean_ctor_set(x_43, 6, x_41); +lean_ctor_set(x_43, 7, x_42); +x_44 = lean_unbox(x_33); +lean_dec(x_33); +lean_ctor_set_uint8(x_43, sizeof(void*)*8, x_44); +lean_ctor_set(x_31, 0, x_43); +return x_31; } else { -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; -x_50 = lean_ctor_get(x_43, 0); -x_51 = lean_ctor_get(x_43, 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; uint8_t x_57; lean_object* x_58; +x_45 = lean_ctor_get(x_31, 0); +x_46 = lean_ctor_get(x_31, 1); +lean_inc(x_46); +lean_inc(x_45); +lean_dec(x_31); +x_47 = lean_ctor_get(x_1, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_1, 1); +lean_inc(x_48); +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_28); +lean_ctor_set(x_49, 1, x_11); +x_50 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_50, 0, x_49); +x_51 = lean_ctor_get(x_1, 3); lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_43); -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_40); -lean_ctor_set(x_52, 1, x_23); -x_53 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_53, 0, x_52); -if (lean_is_scalar(x_18)) { - x_54 = lean_alloc_ctor(0, 8, 1); -} else { - x_54 = x_18; -} -lean_ctor_set(x_54, 0, x_9); -lean_ctor_set(x_54, 1, x_10); -lean_ctor_set(x_54, 2, x_11); -lean_ctor_set(x_54, 3, x_53); -lean_ctor_set(x_54, 4, x_14); -lean_ctor_set(x_54, 5, x_15); -lean_ctor_set(x_54, 6, x_16); -lean_ctor_set(x_54, 7, x_17); -x_55 = lean_unbox(x_50); -lean_dec(x_50); -lean_ctor_set_uint8(x_54, sizeof(void*)*8, x_55); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set(x_56, 1, x_51); -return x_56; -} -} -else -{ -uint8_t x_57; -lean_dec(x_40); -lean_dec(x_23); -lean_dec(x_18); -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_9); -x_57 = !lean_is_exclusive(x_43); -if (x_57 == 0) -{ -return x_43; -} -else -{ -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_43, 0); -x_59 = lean_ctor_get(x_43, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_43); -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_61; -lean_dec(x_23); -lean_dec(x_18); -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_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_61 = !lean_is_exclusive(x_39); -if (x_61 == 0) -{ -return x_39; -} -else -{ -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); -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; -} +x_52 = lean_ctor_get(x_1, 4); +lean_inc(x_52); +x_53 = lean_ctor_get(x_1, 5); +lean_inc(x_53); +x_54 = lean_ctor_get(x_1, 6); +lean_inc(x_54); +x_55 = lean_ctor_get(x_1, 7); +lean_inc(x_55); +lean_dec(x_1); +x_56 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_56, 0, x_47); +lean_ctor_set(x_56, 1, x_48); +lean_ctor_set(x_56, 2, x_50); +lean_ctor_set(x_56, 3, x_51); +lean_ctor_set(x_56, 4, x_52); +lean_ctor_set(x_56, 5, x_53); +lean_ctor_set(x_56, 6, x_54); +lean_ctor_set(x_56, 7, x_55); +x_57 = lean_unbox(x_45); +lean_dec(x_45); +lean_ctor_set_uint8(x_56, sizeof(void*)*8, x_57); +x_58 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_58, 0, x_56); +lean_ctor_set(x_58, 1, x_46); +return x_58; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; -x_65 = lean_ctor_get(x_28, 1); -lean_inc(x_65); +uint8_t x_59; lean_dec(x_28); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_29); -lean_ctor_set(x_66, 1, x_65); -x_67 = lean_alloc_ctor(0, 5, 1); -lean_ctor_set(x_67, 0, x_31); -lean_ctor_set(x_67, 1, x_32); -lean_ctor_set(x_67, 2, x_33); -lean_ctor_set(x_67, 3, x_66); -lean_ctor_set(x_67, 4, x_35); -lean_ctor_set_uint8(x_67, sizeof(void*)*5, x_34); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_68 = l_Lean_PrettyPrinter_Delaborator_delab(x_67, x_3, x_4, x_5, x_6, x_7, x_30); -if (lean_obj_tag(x_68) == 0) -{ -lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; -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_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4; -x_72 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_70); -if (lean_obj_tag(x_72) == 0) -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_72, 1); -lean_inc(x_74); -if (lean_is_exclusive(x_72)) { - lean_ctor_release(x_72, 0); - lean_ctor_release(x_72, 1); - x_75 = x_72; -} else { - lean_dec_ref(x_72); - x_75 = lean_box(0); -} -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_69); -lean_ctor_set(x_76, 1, x_23); -x_77 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_77, 0, x_76); -if (lean_is_scalar(x_18)) { - x_78 = lean_alloc_ctor(0, 8, 1); -} else { - x_78 = x_18; -} -lean_ctor_set(x_78, 0, x_9); -lean_ctor_set(x_78, 1, x_10); -lean_ctor_set(x_78, 2, x_11); -lean_ctor_set(x_78, 3, x_77); -lean_ctor_set(x_78, 4, x_14); -lean_ctor_set(x_78, 5, x_15); -lean_ctor_set(x_78, 6, x_16); -lean_ctor_set(x_78, 7, x_17); -x_79 = lean_unbox(x_73); -lean_dec(x_73); -lean_ctor_set_uint8(x_78, sizeof(void*)*8, x_79); -if (lean_is_scalar(x_75)) { - x_80 = lean_alloc_ctor(0, 2, 0); -} else { - x_80 = x_75; -} -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_74); -return x_80; -} -else -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; -lean_dec(x_69); -lean_dec(x_23); -lean_dec(x_18); -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_9); -x_81 = lean_ctor_get(x_72, 0); -lean_inc(x_81); -x_82 = lean_ctor_get(x_72, 1); -lean_inc(x_82); -if (lean_is_exclusive(x_72)) { - lean_ctor_release(x_72, 0); - lean_ctor_release(x_72, 1); - x_83 = x_72; -} else { - lean_dec_ref(x_72); - x_83 = lean_box(0); -} -if (lean_is_scalar(x_83)) { - x_84 = lean_alloc_ctor(1, 2, 0); -} else { - x_84 = x_83; -} -lean_ctor_set(x_84, 0, x_81); -lean_ctor_set(x_84, 1, x_82); -return x_84; -} +lean_dec(x_1); +x_59 = !lean_is_exclusive(x_31); +if (x_59 == 0) +{ +return x_31; } else { -lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -lean_dec(x_23); -lean_dec(x_18); -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_9); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_85 = lean_ctor_get(x_68, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_68, 1); -lean_inc(x_86); -if (lean_is_exclusive(x_68)) { - lean_ctor_release(x_68, 0); - lean_ctor_release(x_68, 1); - x_87 = x_68; -} else { - lean_dec_ref(x_68); - x_87 = lean_box(0); -} -if (lean_is_scalar(x_87)) { - x_88 = lean_alloc_ctor(1, 2, 0); -} else { - x_88 = x_87; -} -lean_ctor_set(x_88, 0, x_85); -lean_ctor_set(x_88, 1, x_86); -return x_88; +lean_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_31, 0); +x_61 = lean_ctor_get(x_31, 1); +lean_inc(x_61); +lean_inc(x_60); +lean_dec(x_31); +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_89; -lean_dec(x_23); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); +uint8_t x_63; lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_89 = !lean_is_exclusive(x_27); -if (x_89 == 0) +lean_dec(x_1); +x_63 = !lean_is_exclusive(x_27); +if (x_63 == 0) { return x_27; } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; -x_90 = lean_ctor_get(x_27, 0); -x_91 = lean_ctor_get(x_27, 1); -lean_inc(x_91); -lean_inc(x_90); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_27, 0); +x_65 = lean_ctor_get(x_27, 1); +lean_inc(x_65); +lean_inc(x_64); lean_dec(x_27); -x_92 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_92, 0, x_90); -lean_ctor_set(x_92, 1, x_91); -return x_92; -} -} -} -else -{ -lean_object* x_93; lean_object* x_94; uint8_t x_95; -x_93 = lean_array_get_size(x_14); -x_94 = lean_ctor_get(x_9, 1); -lean_inc(x_94); -x_95 = lean_nat_dec_lt(x_93, x_94); -lean_dec(x_94); -if (x_95 == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; -lean_dec(x_93); -x_96 = lean_array_get_size(x_16); -x_97 = lean_ctor_get(x_9, 2); -lean_inc(x_97); -x_98 = lean_array_get_size(x_97); -x_99 = lean_nat_dec_lt(x_96, x_98); -lean_dec(x_98); -if (x_99 == 0) -{ -lean_object* x_100; -lean_dec(x_97); -lean_dec(x_96); -x_100 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_100) == 0) -{ -uint8_t x_101; -x_101 = !lean_is_exclusive(x_100); -if (x_101 == 0) -{ -lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_102 = lean_ctor_get(x_100, 0); -x_103 = lean_array_push(x_17, x_102); -if (lean_is_scalar(x_18)) { - x_104 = lean_alloc_ctor(0, 8, 1); -} else { - x_104 = x_18; -} -lean_ctor_set(x_104, 0, x_9); -lean_ctor_set(x_104, 1, x_10); -lean_ctor_set(x_104, 2, x_11); -lean_ctor_set(x_104, 3, x_12); -lean_ctor_set(x_104, 4, x_14); -lean_ctor_set(x_104, 5, x_15); -lean_ctor_set(x_104, 6, x_16); -lean_ctor_set(x_104, 7, x_103); -lean_ctor_set_uint8(x_104, sizeof(void*)*8, x_13); -lean_ctor_set(x_100, 0, x_104); -return x_100; -} -else -{ -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_105 = lean_ctor_get(x_100, 0); -x_106 = lean_ctor_get(x_100, 1); -lean_inc(x_106); -lean_inc(x_105); -lean_dec(x_100); -x_107 = lean_array_push(x_17, x_105); -if (lean_is_scalar(x_18)) { - x_108 = lean_alloc_ctor(0, 8, 1); -} else { - x_108 = x_18; +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_ctor_set(x_108, 0, x_9); -lean_ctor_set(x_108, 1, x_10); -lean_ctor_set(x_108, 2, x_11); -lean_ctor_set(x_108, 3, x_12); -lean_ctor_set(x_108, 4, x_14); -lean_ctor_set(x_108, 5, x_15); -lean_ctor_set(x_108, 6, x_16); -lean_ctor_set(x_108, 7, x_107); -lean_ctor_set_uint8(x_108, sizeof(void*)*8, x_13); -x_109 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_109, 0, x_108); -lean_ctor_set(x_109, 1, x_106); -return x_109; } } else { -uint8_t x_110; -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_12); +uint8_t x_67; lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_110 = !lean_is_exclusive(x_100); -if (x_110 == 0) +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_67 = !lean_is_exclusive(x_15); +if (x_67 == 0) { -return x_100; +return x_15; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_100, 0); -x_112 = lean_ctor_get(x_100, 1); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_100); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +lean_object* x_68; lean_object* x_69; lean_object* x_70; +x_68 = lean_ctor_get(x_15, 0); +x_69 = lean_ctor_get(x_15, 1); +lean_inc(x_69); +lean_inc(x_68); +lean_dec(x_15); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_68); +lean_ctor_set(x_70, 1, x_69); +return x_70; } } } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_array_fget(x_97, x_96); -lean_dec(x_96); -lean_dec(x_97); -x_115 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2; -x_116 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders___rarg(x_114, x_115, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_116) == 0) +lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; uint8_t x_75; +x_71 = lean_ctor_get(x_1, 3); +lean_inc(x_71); +x_72 = lean_array_get_size(x_71); +x_73 = lean_ctor_get(x_1, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_73, 1); +lean_inc(x_74); +x_75 = lean_nat_dec_lt(x_72, x_74); +lean_dec(x_74); +lean_dec(x_72); +if (x_75 == 0) { -uint8_t x_117; -x_117 = !lean_is_exclusive(x_116); -if (x_117 == 0) +lean_object* x_76; lean_object* x_77; lean_object* x_78; uint8_t x_79; +x_76 = lean_ctor_get(x_1, 5); +lean_inc(x_76); +x_77 = lean_array_get_size(x_76); +x_78 = l_Lean_Meta_Match_MatcherInfo_numAlts(x_73); +x_79 = lean_nat_dec_lt(x_77, x_78); +lean_dec(x_78); +lean_dec(x_77); +if (x_79 == 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; -x_118 = lean_ctor_get(x_116, 0); -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 = lean_array_push(x_15, x_119); -x_122 = lean_array_push(x_16, x_120); -if (lean_is_scalar(x_18)) { - x_123 = lean_alloc_ctor(0, 8, 1); -} else { - x_123 = x_18; -} -lean_ctor_set(x_123, 0, x_9); -lean_ctor_set(x_123, 1, x_10); -lean_ctor_set(x_123, 2, x_11); -lean_ctor_set(x_123, 3, x_12); -lean_ctor_set(x_123, 4, x_14); -lean_ctor_set(x_123, 5, x_121); -lean_ctor_set(x_123, 6, x_122); -lean_ctor_set(x_123, 7, x_17); -lean_ctor_set_uint8(x_123, sizeof(void*)*8, x_13); -lean_ctor_set(x_116, 0, x_123); -return x_116; +lean_object* x_80; lean_object* x_81; +lean_dec(x_76); +lean_dec(x_73); +lean_dec(x_71); +lean_dec(x_9); +lean_dec(x_1); +x_80 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__4; +x_81 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5(x_80, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +return x_81; } 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; -x_124 = lean_ctor_get(x_116, 0); -x_125 = lean_ctor_get(x_116, 1); -lean_inc(x_125); -lean_inc(x_124); -lean_dec(x_116); -x_126 = lean_ctor_get(x_124, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_124, 1); -lean_inc(x_127); -lean_dec(x_124); -x_128 = lean_array_push(x_15, x_126); -x_129 = lean_array_push(x_16, x_127); -if (lean_is_scalar(x_18)) { - x_130 = lean_alloc_ctor(0, 8, 1); -} else { - x_130 = x_18; -} -lean_ctor_set(x_130, 0, x_9); -lean_ctor_set(x_130, 1, x_10); -lean_ctor_set(x_130, 2, x_11); -lean_ctor_set(x_130, 3, x_12); -lean_ctor_set(x_130, 4, x_14); -lean_ctor_set(x_130, 5, x_128); -lean_ctor_set(x_130, 6, x_129); -lean_ctor_set(x_130, 7, x_17); -lean_ctor_set_uint8(x_130, sizeof(void*)*8, x_13); -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_125); -return x_131; +lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_82 = lean_ctor_get(x_2, 3); +lean_inc(x_82); +lean_dec(x_2); +x_83 = lean_ctor_get(x_1, 1); +lean_inc(x_83); +x_84 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_85 = lean_ctor_get(x_1, 4); +lean_inc(x_85); +x_86 = lean_array_push(x_76, x_82); +x_87 = lean_ctor_get(x_1, 6); +lean_inc(x_87); +x_88 = lean_ctor_get(x_1, 7); +lean_inc(x_88); +lean_dec(x_1); +x_89 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_89, 0, x_73); +lean_ctor_set(x_89, 1, x_83); +lean_ctor_set(x_89, 2, x_9); +lean_ctor_set(x_89, 3, x_71); +lean_ctor_set(x_89, 4, x_85); +lean_ctor_set(x_89, 5, x_86); +lean_ctor_set(x_89, 6, x_87); +lean_ctor_set(x_89, 7, x_88); +lean_ctor_set_uint8(x_89, sizeof(void*)*8, x_84); +x_90 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_90, 0, x_89); +lean_ctor_set(x_90, 1, x_8); +return x_90; } } else { -uint8_t x_132; -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -x_132 = !lean_is_exclusive(x_116); -if (x_132 == 0) +lean_object* x_91; +x_91 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_91) == 0) { -return x_116; +uint8_t x_92; +x_92 = !lean_is_exclusive(x_91); +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; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_93 = lean_ctor_get(x_91, 0); +x_94 = lean_ctor_get(x_1, 1); +lean_inc(x_94); +x_95 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_96 = lean_array_push(x_71, x_93); +x_97 = lean_ctor_get(x_1, 4); +lean_inc(x_97); +x_98 = lean_ctor_get(x_1, 5); +lean_inc(x_98); +x_99 = lean_ctor_get(x_1, 6); +lean_inc(x_99); +x_100 = lean_ctor_get(x_1, 7); +lean_inc(x_100); +lean_dec(x_1); +x_101 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_101, 0, x_73); +lean_ctor_set(x_101, 1, x_94); +lean_ctor_set(x_101, 2, x_9); +lean_ctor_set(x_101, 3, x_96); +lean_ctor_set(x_101, 4, x_97); +lean_ctor_set(x_101, 5, x_98); +lean_ctor_set(x_101, 6, x_99); +lean_ctor_set(x_101, 7, x_100); +lean_ctor_set_uint8(x_101, sizeof(void*)*8, x_95); +lean_ctor_set(x_91, 0, x_101); +return x_91; } else { -lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_116, 0); -x_134 = lean_ctor_get(x_116, 1); -lean_inc(x_134); -lean_inc(x_133); -lean_dec(x_116); -x_135 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -return x_135; -} -} +lean_object* x_102; lean_object* x_103; lean_object* x_104; uint8_t x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_102 = lean_ctor_get(x_91, 0); +x_103 = lean_ctor_get(x_91, 1); +lean_inc(x_103); +lean_inc(x_102); +lean_dec(x_91); +x_104 = lean_ctor_get(x_1, 1); +lean_inc(x_104); +x_105 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_106 = lean_array_push(x_71, x_102); +x_107 = lean_ctor_get(x_1, 4); +lean_inc(x_107); +x_108 = lean_ctor_get(x_1, 5); +lean_inc(x_108); +x_109 = lean_ctor_get(x_1, 6); +lean_inc(x_109); +x_110 = lean_ctor_get(x_1, 7); +lean_inc(x_110); +lean_dec(x_1); +x_111 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_111, 0, x_73); +lean_ctor_set(x_111, 1, x_104); +lean_ctor_set(x_111, 2, x_9); +lean_ctor_set(x_111, 3, x_106); +lean_ctor_set(x_111, 4, x_107); +lean_ctor_set(x_111, 5, x_108); +lean_ctor_set(x_111, 6, x_109); +lean_ctor_set(x_111, 7, x_110); +lean_ctor_set_uint8(x_111, sizeof(void*)*8, x_105); +x_112 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_112, 0, x_111); +lean_ctor_set(x_112, 1, x_103); +return x_112; } } else { -lean_object* x_136; -lean_inc(x_6); -x_136 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_136) == 0) -{ -lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_167; lean_object* x_168; uint8_t x_169; -x_137 = lean_ctor_get(x_136, 0); -lean_inc(x_137); -x_138 = lean_ctor_get(x_136, 1); -lean_inc(x_138); -if (lean_is_exclusive(x_136)) { - lean_ctor_release(x_136, 0); - lean_ctor_release(x_136, 1); - x_139 = x_136; -} else { - lean_dec_ref(x_136); - x_139 = lean_box(0); -} -x_167 = lean_ctor_get(x_9, 4); -lean_inc(x_167); -x_168 = lean_array_get_size(x_167); -x_169 = lean_nat_dec_lt(x_93, x_168); -lean_dec(x_168); -if (x_169 == 0) +uint8_t x_113; +lean_dec(x_73); +lean_dec(x_71); +lean_dec(x_9); +lean_dec(x_1); +x_113 = !lean_is_exclusive(x_91); +if (x_113 == 0) { -lean_object* x_170; lean_object* x_171; -lean_dec(x_167); -lean_dec(x_93); -x_170 = l_Lean_Meta_Match_instInhabitedDiscrInfo; -x_171 = l___private_Init_GetElem_0__outOfBounds___rarg(x_170); -x_140 = x_171; -goto block_166; +return x_91; } else { -lean_object* x_172; -x_172 = lean_array_fget(x_167, x_93); -lean_dec(x_93); -lean_dec(x_167); -x_140 = x_172; -goto block_166; -} -block_166: -{ -if (lean_obj_tag(x_140) == 0) -{ -lean_object* x_141; uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_141 = lean_ctor_get(x_6, 5); -lean_inc(x_141); -lean_dec(x_6); -x_142 = 0; -x_143 = l_Lean_SourceInfo_fromRef(x_141, x_142); -x_144 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_145 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_143); -x_146 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_146, 0, x_143); -lean_ctor_set(x_146, 1, x_144); -lean_ctor_set(x_146, 2, x_145); -x_147 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4; -x_148 = l_Lean_Syntax_node2(x_143, x_147, x_146, x_137); -x_149 = lean_array_push(x_14, x_148); -if (lean_is_scalar(x_18)) { - x_150 = lean_alloc_ctor(0, 8, 1); -} else { - x_150 = x_18; -} -lean_ctor_set(x_150, 0, x_9); -lean_ctor_set(x_150, 1, x_10); -lean_ctor_set(x_150, 2, x_11); -lean_ctor_set(x_150, 3, x_12); -lean_ctor_set(x_150, 4, x_149); -lean_ctor_set(x_150, 5, x_15); -lean_ctor_set(x_150, 6, x_16); -lean_ctor_set(x_150, 7, x_17); -lean_ctor_set_uint8(x_150, sizeof(void*)*8, x_13); -if (lean_is_scalar(x_139)) { - x_151 = lean_alloc_ctor(0, 2, 0); -} else { - x_151 = x_139; +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; } -lean_ctor_set(x_151, 0, x_150); -lean_ctor_set(x_151, 1, x_138); -return x_151; } -else -{ -lean_object* x_152; lean_object* x_153; uint8_t x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_152 = lean_ctor_get(x_140, 0); -lean_inc(x_152); -lean_dec(x_140); -x_153 = lean_ctor_get(x_6, 5); -lean_inc(x_153); -lean_dec(x_6); -x_154 = 0; -x_155 = l_Lean_SourceInfo_fromRef(x_153, x_154); -x_156 = lean_mk_syntax_ident(x_152); -x_157 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__8; -lean_inc(x_155); -x_158 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_158, 0, x_155); -lean_ctor_set(x_158, 1, x_157); -x_159 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_155); -x_160 = l_Lean_Syntax_node2(x_155, x_159, x_156, x_158); -x_161 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4; -x_162 = l_Lean_Syntax_node2(x_155, x_161, x_160, x_137); -x_163 = lean_array_push(x_14, x_162); -if (lean_is_scalar(x_18)) { - x_164 = lean_alloc_ctor(0, 8, 1); -} else { - x_164 = x_18; -} -lean_ctor_set(x_164, 0, x_9); -lean_ctor_set(x_164, 1, x_10); -lean_ctor_set(x_164, 2, x_11); -lean_ctor_set(x_164, 3, x_12); -lean_ctor_set(x_164, 4, x_163); -lean_ctor_set(x_164, 5, x_15); -lean_ctor_set(x_164, 6, x_16); -lean_ctor_set(x_164, 7, x_17); -lean_ctor_set_uint8(x_164, sizeof(void*)*8, x_13); -if (lean_is_scalar(x_139)) { - x_165 = lean_alloc_ctor(0, 2, 0); -} else { - x_165 = x_139; } -lean_ctor_set(x_165, 0, x_164); -lean_ctor_set(x_165, 1, x_138); -return x_165; } } } -else +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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: { -uint8_t x_173; -lean_dec(x_93); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -x_173 = !lean_is_exclusive(x_136); -if (x_173 == 0) +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_11 = lean_array_get_size(x_3); +x_12 = lean_unsigned_to_nat(0u); +x_13 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_3, x_12, x_11); +x_14 = lean_ctor_get(x_1, 1); +x_15 = lean_ctor_get(x_1, 2); +x_16 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); +x_17 = lean_ctor_get(x_1, 3); +x_18 = lean_ctor_get(x_1, 5); +x_19 = lean_ctor_get(x_1, 6); +x_20 = lean_ctor_get(x_1, 7); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_15); +lean_inc(x_14); +x_21 = lean_alloc_ctor(0, 8, 1); +lean_ctor_set(x_21, 0, x_2); +lean_ctor_set(x_21, 1, x_14); +lean_ctor_set(x_21, 2, x_15); +lean_ctor_set(x_21, 3, x_17); +lean_ctor_set(x_21, 4, x_3); +lean_ctor_set(x_21, 5, x_18); +lean_ctor_set(x_21, 6, x_19); +lean_ctor_set(x_21, 7, x_20); +lean_ctor_set_uint8(x_21, sizeof(void*)*8, x_16); +x_22 = lean_array_get_size(x_18); +x_23 = lean_unsigned_to_nat(1u); +lean_inc(x_22); +x_24 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_13, x_22, x_12, x_22, x_23, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_22); +if (lean_obj_tag(x_24) == 0) { -return x_136; +uint8_t x_25; +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) +{ +return x_24; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -x_174 = lean_ctor_get(x_136, 0); -x_175 = lean_ctor_get(x_136, 1); -lean_inc(x_175); -lean_inc(x_174); -lean_dec(x_136); -x_176 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_176, 0, x_174); -lean_ctor_set(x_176, 1, x_175); -return x_176; -} -} -} +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_24, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_24); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_26); +lean_ctor_set(x_28, 1, x_27); +return x_28; } } else { -lean_object* x_177; uint8_t x_178; -x_177 = 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); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -x_178 = !lean_is_exclusive(x_177); -if (x_178 == 0) +uint8_t x_29; +x_29 = !lean_is_exclusive(x_24); +if (x_29 == 0) { -lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_179 = lean_ctor_get(x_177, 0); -x_180 = lean_array_push(x_11, x_179); -if (lean_is_scalar(x_18)) { - x_181 = lean_alloc_ctor(0, 8, 1); -} else { - x_181 = x_18; -} -lean_ctor_set(x_181, 0, x_9); -lean_ctor_set(x_181, 1, x_10); -lean_ctor_set(x_181, 2, x_180); -lean_ctor_set(x_181, 3, x_12); -lean_ctor_set(x_181, 4, x_14); -lean_ctor_set(x_181, 5, x_15); -lean_ctor_set(x_181, 6, x_16); -lean_ctor_set(x_181, 7, x_17); -lean_ctor_set_uint8(x_181, sizeof(void*)*8, x_13); -lean_ctor_set(x_177, 0, x_181); -return x_177; +return x_24; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_182 = lean_ctor_get(x_177, 0); -x_183 = lean_ctor_get(x_177, 1); -lean_inc(x_183); -lean_inc(x_182); -lean_dec(x_177); -x_184 = lean_array_push(x_11, x_182); -if (lean_is_scalar(x_18)) { - x_185 = lean_alloc_ctor(0, 8, 1); -} else { - x_185 = x_18; -} -lean_ctor_set(x_185, 0, x_9); -lean_ctor_set(x_185, 1, x_10); -lean_ctor_set(x_185, 2, x_184); -lean_ctor_set(x_185, 3, x_12); -lean_ctor_set(x_185, 4, x_14); -lean_ctor_set(x_185, 5, x_15); -lean_ctor_set(x_185, 6, x_16); -lean_ctor_set(x_185, 7, x_17); -lean_ctor_set_uint8(x_185, sizeof(void*)*8, x_13); -x_186 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_186, 0, x_185); -lean_ctor_set(x_186, 1, x_183); -return x_186; +lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_24, 0); +x_31 = lean_ctor_get(x_24, 1); +lean_inc(x_31); +lean_inc(x_30); +lean_dec(x_24); +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_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_ctor_get(x_1, 7); -lean_inc(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_dec(x_3); +x_11 = lean_ctor_get(x_1, 3); +lean_inc(x_11); +x_12 = lean_array_get_size(x_11); +lean_dec(x_11); +x_13 = lean_unsigned_to_nat(1u); +x_14 = lean_nat_add(x_13, x_12); +lean_dec(x_12); +x_15 = lean_array_get_size(x_2); +lean_inc(x_14); +x_16 = l_Array_toSubarray___rarg(x_2, x_14, x_15); +x_17 = l_Array_ofSubarray___rarg(x_16); +x_18 = lean_array_get_size(x_17); +x_19 = lean_mk_empty_array_with_capacity(x_18); +x_20 = lean_unsigned_to_nat(0u); +x_21 = l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18(x_1, x_14, x_17, x_17, x_18, x_20, lean_box(0), x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_17); +lean_dec(x_14); lean_dec(x_1); -x_11 = l_Lean_Syntax_mkApp(x_2, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; +return x_21; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("term", 4); +x_1 = lean_mk_string_from_bytes("match", 5); return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; -x_3 = l_Lean_Name_str___override(x_1, x_2); -return x_3; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +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_delabAppMatch___lambda__6___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_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3() { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; -x_3 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_3, 0, x_2); -lean_ctor_set(x_3, 1, x_1); -return x_3; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("with", 4); +return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("match", 5); +x_1 = lean_mk_string_from_bytes("matchAlts", 9); return x_1; } } @@ -24821,880 +24660,1068 @@ return x_5; static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("with", 4); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_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_delabAppImplicitCore_mkArg___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_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__7() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("matchAlts", 9); +x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___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_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_delabAppMatch___lambda__6___closed__7; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Option.get!", 11); +return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___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_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_delabAppImplicitCore_mkArgStx___closed__5; -x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); -return x_5; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("value is none", 13); +return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10() { _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_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__7; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__9; +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_delabAppMatch___lambda__6___closed__11() { +_start: +{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("nomatch", 7); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___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_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_delabAppMatch___lambda__6___closed__10; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; 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_delabAppMatch___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_dec(x_2); -x_10 = lean_ctor_get(x_1, 4); +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; size_t x_15; size_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -x_11 = lean_ctor_get(x_1, 6); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -x_12 = lean_array_get_size(x_10); -x_13 = lean_unsigned_to_nat(1u); -x_14 = lean_nat_dec_eq(x_12, x_13); -lean_dec(x_12); -if (x_14 == 0) +lean_dec(x_9); +x_12 = lean_ctor_get(x_1, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_12, 4); +lean_inc(x_13); +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_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_15, x_16, x_13); +x_18 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___boxed), 10, 2); +lean_closure_set(x_18, 0, x_1); +lean_closure_set(x_18, 1, x_12); +x_19 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +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_20 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch_withDummyBinders___rarg(x_17, x_10, x_18, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_11); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_15 = lean_array_get_size(x_11); -x_16 = lean_unsigned_to_nat(0u); -x_17 = lean_nat_dec_eq(x_15, x_16); -lean_dec(x_15); -if (x_17 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_18; -lean_inc(x_8); +lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_ctor_get(x_20, 1); +x_24 = lean_ctor_get(x_22, 7); +lean_inc(x_24); +x_25 = l_Array_isEmpty___rarg(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; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_free_object(x_20); +x_26 = lean_ctor_get(x_2, 0); +lean_inc(x_26); +x_27 = lean_ctor_get(x_2, 1); +lean_inc(x_27); +x_28 = lean_ctor_get(x_2, 2); +lean_inc(x_28); +x_29 = lean_ctor_get(x_22, 1); +lean_inc(x_29); +x_30 = lean_ctor_get(x_2, 4); +lean_inc(x_30); +x_31 = 1; +x_32 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_32, 0, x_26); +lean_ctor_set(x_32, 1, x_27); +lean_ctor_set(x_32, 2, x_28); +lean_ctor_set(x_32, 3, x_29); +lean_ctor_set(x_32, 4, x_30); +lean_ctor_set_uint8(x_32, sizeof(void*)*5, x_31); +x_33 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_32, x_3, x_4, x_5, x_6, x_7, x_23); +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); +lean_inc(x_22); +x_36 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5), 10, 1); +lean_closure_set(x_36, 0, x_22); 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_18 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_18) == 0) +x_37 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg(x_34, x_36, x_32, x_3, x_4, x_5, x_6, x_7, x_35); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_108; -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_108 = lean_ctor_get(x_1, 3); -lean_inc(x_108); -if (lean_obj_tag(x_108) == 0) +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; size_t x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_132; +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_22, 4); +lean_inc(x_40); +x_41 = lean_ctor_get(x_22, 3); +lean_inc(x_41); +x_42 = l_Array_zip___rarg(x_40, x_41); +lean_dec(x_41); +lean_dec(x_40); +x_43 = lean_array_get_size(x_42); +x_44 = lean_usize_of_nat(x_43); +lean_dec(x_43); +lean_inc(x_6); +x_45 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19(x_44, x_16, x_42, x_2, x_3, x_4, x_5, x_6, x_7, x_39); +lean_dec(x_3); +lean_dec(x_2); +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); +lean_inc(x_47); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_48 = x_45; +} else { + lean_dec_ref(x_45); + x_48 = lean_box(0); +} +x_132 = lean_ctor_get(x_22, 2); +lean_inc(x_132); +if (lean_obj_tag(x_132) == 0) { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4; -x_110 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_109); -x_111 = lean_ctor_get(x_110, 0); -lean_inc(x_111); -x_112 = lean_ctor_get(x_110, 1); -lean_inc(x_112); -lean_dec(x_110); -x_21 = x_111; -x_22 = x_112; -goto block_107; +lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; +x_133 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10; +x_134 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23(x_133); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_49 = x_135; +x_50 = x_136; +goto block_131; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_108, 0); -lean_inc(x_113); -lean_dec(x_108); -x_114 = lean_ctor_get(x_113, 0); -lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -lean_dec(x_113); -x_21 = x_114; -x_22 = x_115; -goto block_107; +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_132, 0); +lean_inc(x_137); +lean_dec(x_132); +x_138 = lean_ctor_get(x_137, 0); +lean_inc(x_138); +x_139 = lean_ctor_get(x_137, 1); +lean_inc(x_139); +lean_dec(x_137); +x_49 = x_138; +x_50 = x_139; +goto block_131; } -block_107: +block_131: { -lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; -x_23 = lean_ctor_get(x_7, 2); -lean_inc(x_23); -x_24 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -if (x_24 == 0) +lean_object* x_51; uint8_t x_52; uint8_t x_53; lean_object* x_54; +x_51 = lean_ctor_get(x_6, 2); +lean_inc(x_51); +x_52 = lean_ctor_get_uint8(x_22, sizeof(void*)*8); +lean_dec(x_22); +if (x_52 == 0) { -lean_object* x_99; -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_123; lean_inc(x_6); -lean_inc(x_5); -x_99 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_22, x_23, x_5, x_6, x_7, x_8, x_20); -lean_dec(x_23); -if (lean_obj_tag(x_99) == 0) +x_123 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_50, x_51, x_4, x_5, x_6, x_7, x_47); +lean_dec(x_51); +if (lean_obj_tag(x_123) == 0) { -lean_object* x_100; lean_object* x_101; uint8_t x_102; -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_unbox(x_100); -lean_dec(x_100); -x_25 = x_102; -x_26 = x_101; -goto block_98; +lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +lean_dec(x_123); +x_126 = lean_unbox(x_124); +lean_dec(x_124); +x_53 = x_126; +x_54 = x_125; +goto block_122; } else { -uint8_t x_103; -lean_dec(x_21); -lean_dec(x_19); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); +uint8_t x_127; +lean_dec(x_49); +lean_dec(x_48); +lean_dec(x_46); +lean_dec(x_38); +lean_dec(x_24); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_103 = !lean_is_exclusive(x_99); -if (x_103 == 0) +x_127 = !lean_is_exclusive(x_123); +if (x_127 == 0) { -return x_99; +return x_123; } else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; -x_104 = lean_ctor_get(x_99, 0); -x_105 = lean_ctor_get(x_99, 1); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_99); -x_106 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_106, 0, x_104); -lean_ctor_set(x_106, 1, x_105); -return x_106; +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_123, 0); +x_129 = lean_ctor_get(x_123, 1); +lean_inc(x_129); +lean_inc(x_128); +lean_dec(x_123); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +return x_130; } } } else { -lean_dec(x_23); -lean_dec(x_22); -x_25 = x_24; -x_26 = x_20; -goto block_98; -} -block_98: -{ -if (x_25 == 0) -{ -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; lean_object* x_34; lean_object* x_35; size_t x_36; size_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; size_t x_47; lean_object* x_48; lean_object* x_49; 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_dec(x_21); -x_27 = lean_ctor_get(x_7, 5); -lean_inc(x_27); -x_28 = 0; -x_29 = l_Lean_SourceInfo_fromRef(x_27, x_28); -x_30 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4; -lean_inc(x_29); -x_31 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_31, 0, x_29); -lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_33 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_29); -x_34 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_34, 0, x_29); -lean_ctor_set(x_34, 1, x_32); -lean_ctor_set(x_34, 2, x_33); -x_35 = lean_array_get_size(x_10); -x_36 = lean_usize_of_nat(x_35); -lean_dec(x_35); -x_37 = 0; -x_38 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_36, x_37, x_10); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; -x_40 = l_Lean_mkSepArray(x_38, x_39); -lean_dec(x_38); -x_41 = l_Array_append___rarg(x_33, x_40); -lean_inc(x_29); -x_42 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_42, 0, x_29); -lean_ctor_set(x_42, 1, x_32); -lean_ctor_set(x_42, 2, x_41); -x_43 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; -lean_inc(x_29); -x_44 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_44, 0, x_29); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Array_zip___rarg(x_19, x_11); -lean_dec(x_11); -lean_dec(x_19); -x_46 = lean_array_get_size(x_45); -x_47 = lean_usize_of_nat(x_46); -lean_dec(x_46); -x_48 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; -lean_inc(x_29); -x_49 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_48, x_29, x_32, x_33, x_47, x_37, x_45); -x_50 = l_Array_append___rarg(x_33, x_49); -lean_inc(x_29); -x_51 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_51, 0, x_29); -lean_ctor_set(x_51, 1, x_32); -lean_ctor_set(x_51, 2, x_50); -x_52 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; -lean_inc(x_29); -x_53 = l_Lean_Syntax_node1(x_29, x_52, x_51); -x_54 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; -lean_inc(x_34); -x_55 = l_Lean_Syntax_node6(x_29, x_54, x_31, x_34, x_34, x_42, x_44, x_53); -x_56 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(x_1, x_55, x_3, x_4, x_5, x_6, x_7, x_8, x_26); -lean_dec(x_8); +lean_dec(x_51); +lean_dec(x_50); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -return x_56; +x_53 = x_52; +x_54 = x_47; +goto block_122; } -else +block_122: +{ +if (x_53 == 0) { -lean_object* x_57; uint8_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; size_t x_77; 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; size_t x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_57 = lean_ctor_get(x_7, 5); +lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; 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; 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_dec(x_49); +x_55 = lean_ctor_get(x_6, 5); +lean_inc(x_55); +lean_dec(x_6); +x_56 = 0; +x_57 = l_Lean_SourceInfo_fromRef(x_55, x_56); +x_58 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; lean_inc(x_57); -x_58 = 0; -x_59 = l_Lean_SourceInfo_fromRef(x_57, x_58); -x_60 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4; -lean_inc(x_59); -x_61 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_61, 0, x_59); +x_59 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_58); +x_60 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_57); +x_61 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_61, 0, x_57); lean_ctor_set(x_61, 1, x_60); -x_62 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_63 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_59); -x_64 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_64, 0, x_59); -lean_ctor_set(x_64, 1, x_62); -lean_ctor_set(x_64, 2, x_63); -x_65 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; -lean_inc(x_59); -x_66 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_66, 0, x_59); -lean_ctor_set(x_66, 1, x_65); -x_67 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5; -lean_inc(x_59); -x_68 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_68, 0, x_59); -lean_ctor_set(x_68, 1, x_67); -x_69 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; -lean_inc(x_59); +lean_ctor_set(x_61, 2, x_19); +x_62 = lean_array_get_size(x_46); +x_63 = lean_usize_of_nat(x_62); +lean_dec(x_62); +x_64 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_63, x_16, x_46); +x_65 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_66 = l_Lean_mkSepArray(x_64, x_65); +lean_dec(x_64); +x_67 = l_Array_append___rarg(x_19, x_66); +lean_inc(x_57); +x_68 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_68, 0, x_57); +lean_ctor_set(x_68, 1, x_60); +lean_ctor_set(x_68, 2, x_67); +x_69 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; +lean_inc(x_57); x_70 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_70, 0, x_59); +lean_ctor_set(x_70, 0, x_57); lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; -lean_inc(x_59); -x_72 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_72, 0, x_59); -lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__9; -lean_inc(x_59); -x_74 = l_Lean_Syntax_node5(x_59, x_73, x_66, x_68, x_70, x_21, x_72); -lean_inc(x_59); -x_75 = l_Lean_Syntax_node1(x_59, x_62, x_74); -x_76 = lean_array_get_size(x_10); -x_77 = lean_usize_of_nat(x_76); -lean_dec(x_76); -x_78 = 0; -x_79 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_77, x_78, x_10); -x_80 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; -x_81 = l_Lean_mkSepArray(x_79, x_80); -lean_dec(x_79); -x_82 = l_Array_append___rarg(x_63, x_81); -lean_inc(x_59); -x_83 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_83, 0, x_59); -lean_ctor_set(x_83, 1, x_62); -lean_ctor_set(x_83, 2, x_82); -x_84 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; -lean_inc(x_59); -x_85 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_85, 0, x_59); -lean_ctor_set(x_85, 1, x_84); -x_86 = l_Array_zip___rarg(x_19, x_11); -lean_dec(x_11); -lean_dec(x_19); -x_87 = lean_array_get_size(x_86); -x_88 = lean_usize_of_nat(x_87); -lean_dec(x_87); -x_89 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; -lean_inc(x_59); -x_90 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_89, x_59, x_62, x_63, x_88, x_78, x_86); -x_91 = l_Array_append___rarg(x_63, x_90); -lean_inc(x_59); -x_92 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_92, 0, x_59); -lean_ctor_set(x_92, 1, x_62); -lean_ctor_set(x_92, 2, x_91); -x_93 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; -lean_inc(x_59); -x_94 = l_Lean_Syntax_node1(x_59, x_93, x_92); -x_95 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; -x_96 = l_Lean_Syntax_node6(x_59, x_95, x_61, x_64, x_75, x_83, x_85, x_94); -x_97 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(x_1, x_96, x_3, x_4, x_5, x_6, x_7, x_8, x_26); -lean_dec(x_8); -lean_dec(x_7); +x_71 = l_Array_zip___rarg(x_38, x_24); +lean_dec(x_24); +lean_dec(x_38); +x_72 = lean_array_get_size(x_71); +x_73 = lean_usize_of_nat(x_72); +lean_dec(x_72); +x_74 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +lean_inc(x_57); +x_75 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21(x_74, x_57, x_60, x_19, x_73, x_16, x_71); +x_76 = l_Array_append___rarg(x_19, x_75); +lean_inc(x_57); +x_77 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_77, 0, x_57); +lean_ctor_set(x_77, 1, x_60); +lean_ctor_set(x_77, 2, x_76); +x_78 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; +lean_inc(x_57); +x_79 = l_Lean_Syntax_node1(x_57, x_78, x_77); +x_80 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; +lean_inc(x_61); +x_81 = l_Lean_Syntax_node6(x_57, x_80, x_59, x_61, x_61, x_68, x_70, x_79); +if (lean_is_scalar(x_48)) { + x_82 = lean_alloc_ctor(0, 2, 0); +} else { + x_82 = x_48; +} +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_54); +return x_82; +} +else +{ +lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; size_t x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; size_t x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_83 = lean_ctor_get(x_6, 5); +lean_inc(x_83); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_97; +x_84 = 0; +x_85 = l_Lean_SourceInfo_fromRef(x_83, x_84); +x_86 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; +lean_inc(x_85); +x_87 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_87, 0, x_85); +lean_ctor_set(x_87, 1, x_86); +x_88 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_85); +x_89 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_89, 0, x_85); +lean_ctor_set(x_89, 1, x_88); +lean_ctor_set(x_89, 2, x_19); +x_90 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +lean_inc(x_85); +x_91 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_91, 0, x_85); +lean_ctor_set(x_91, 1, x_90); +x_92 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3; +lean_inc(x_85); +x_93 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_93, 0, x_85); +lean_ctor_set(x_93, 1, x_92); +x_94 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; +lean_inc(x_85); +x_95 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_95, 0, x_85); +lean_ctor_set(x_95, 1, x_94); +x_96 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_85); +x_97 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_97, 0, x_85); +lean_ctor_set(x_97, 1, x_96); +x_98 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; +lean_inc(x_85); +x_99 = l_Lean_Syntax_node5(x_85, x_98, x_91, x_93, x_95, x_49, x_97); +lean_inc(x_85); +x_100 = l_Lean_Syntax_node1(x_85, x_88, x_99); +x_101 = lean_array_get_size(x_46); +x_102 = lean_usize_of_nat(x_101); +lean_dec(x_101); +x_103 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_102, x_16, x_46); +x_104 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_105 = l_Lean_mkSepArray(x_103, x_104); +lean_dec(x_103); +x_106 = l_Array_append___rarg(x_19, x_105); +lean_inc(x_85); +x_107 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_107, 0, x_85); +lean_ctor_set(x_107, 1, x_88); +lean_ctor_set(x_107, 2, x_106); +x_108 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; +lean_inc(x_85); +x_109 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_109, 0, x_85); +lean_ctor_set(x_109, 1, x_108); +x_110 = l_Array_zip___rarg(x_38, x_24); +lean_dec(x_24); +lean_dec(x_38); +x_111 = lean_array_get_size(x_110); +x_112 = lean_usize_of_nat(x_111); +lean_dec(x_111); +x_113 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +lean_inc(x_85); +x_114 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22(x_113, x_85, x_88, x_19, x_112, x_16, x_110); +x_115 = l_Array_append___rarg(x_19, x_114); +lean_inc(x_85); +x_116 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_116, 0, x_85); +lean_ctor_set(x_116, 1, x_88); +lean_ctor_set(x_116, 2, x_115); +x_117 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; +lean_inc(x_85); +x_118 = l_Lean_Syntax_node1(x_85, x_117, x_116); +x_119 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; +x_120 = l_Lean_Syntax_node6(x_85, x_119, x_87, x_89, x_100, x_107, x_109, x_118); +if (lean_is_scalar(x_48)) { + x_121 = lean_alloc_ctor(0, 2, 0); +} else { + x_121 = x_48; +} +lean_ctor_set(x_121, 0, x_120); +lean_ctor_set(x_121, 1, x_54); +return x_121; } } } } else { -uint8_t x_116; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); +uint8_t x_140; +lean_dec(x_24); +lean_dec(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_116 = !lean_is_exclusive(x_18); -if (x_116 == 0) +lean_dec(x_2); +x_140 = !lean_is_exclusive(x_37); +if (x_140 == 0) { -return x_18; +return x_37; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_18, 0); -x_118 = lean_ctor_get(x_18, 1); -lean_inc(x_118); -lean_inc(x_117); -lean_dec(x_18); -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_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_37, 0); +x_142 = lean_ctor_get(x_37, 1); +lean_inc(x_142); +lean_inc(x_141); +lean_dec(x_37); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +return x_143; } } } else { -lean_object* x_120; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); +lean_object* x_144; uint8_t 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_dec(x_24); 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_120 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_9); -return x_120; +lean_dec(x_2); +x_144 = lean_ctor_get(x_6, 5); +lean_inc(x_144); +lean_dec(x_6); +x_145 = 0; +x_146 = l_Lean_SourceInfo_fromRef(x_144, x_145); +x_147 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; +lean_inc(x_146); +x_148 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +x_149 = lean_ctor_get(x_22, 3); +lean_inc(x_149); +lean_dec(x_22); +x_150 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; +x_151 = l_Lean_Syntax_SepArray_ofElems(x_150, x_149); +lean_dec(x_149); +x_152 = l_Array_append___rarg(x_19, x_151); +x_153 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_146); +x_154 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_154, 0, x_146); +lean_ctor_set(x_154, 1, x_153); +lean_ctor_set(x_154, 2, x_152); +x_155 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12; +x_156 = l_Lean_Syntax_node2(x_146, x_155, x_148, x_154); +lean_ctor_set(x_20, 0, x_156); +return x_20; } } else { -lean_object* x_121; lean_object* x_122; uint8_t x_123; -x_121 = lean_array_get_size(x_11); -x_122 = lean_unsigned_to_nat(0u); -x_123 = lean_nat_dec_eq(x_121, x_122); -lean_dec(x_121); -if (x_123 == 0) +lean_object* x_157; lean_object* x_158; lean_object* x_159; uint8_t x_160; +x_157 = lean_ctor_get(x_20, 0); +x_158 = lean_ctor_get(x_20, 1); +lean_inc(x_158); +lean_inc(x_157); +lean_dec(x_20); +x_159 = lean_ctor_get(x_157, 7); +lean_inc(x_159); +x_160 = l_Array_isEmpty___rarg(x_159); +if (x_160 == 0) { -lean_object* x_124; -lean_inc(x_8); +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; uint8_t x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_161 = lean_ctor_get(x_2, 0); +lean_inc(x_161); +x_162 = lean_ctor_get(x_2, 1); +lean_inc(x_162); +x_163 = lean_ctor_get(x_2, 2); +lean_inc(x_163); +x_164 = lean_ctor_get(x_157, 1); +lean_inc(x_164); +x_165 = lean_ctor_get(x_2, 4); +lean_inc(x_165); +x_166 = 1; +x_167 = lean_alloc_ctor(0, 5, 1); +lean_ctor_set(x_167, 0, x_161); +lean_ctor_set(x_167, 1, x_162); +lean_ctor_set(x_167, 2, x_163); +lean_ctor_set(x_167, 3, x_164); +lean_ctor_set(x_167, 4, x_165); +lean_ctor_set_uint8(x_167, sizeof(void*)*5, x_166); +x_168 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_167, x_3, x_4, x_5, x_6, x_7, x_158); +x_169 = lean_ctor_get(x_168, 0); +lean_inc(x_169); +x_170 = lean_ctor_get(x_168, 1); +lean_inc(x_170); +lean_dec(x_168); +lean_inc(x_157); +x_171 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5), 10, 1); +lean_closure_set(x_171, 0, x_157); 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_124 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_124) == 0) +x_172 = l_Lean_Meta_forallTelescopeReducing___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___spec__1___rarg(x_169, x_171, x_167, x_3, x_4, x_5, x_6, x_7, x_170); +if (lean_obj_tag(x_172) == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_214; -x_125 = lean_ctor_get(x_124, 0); -lean_inc(x_125); -x_126 = lean_ctor_get(x_124, 1); -lean_inc(x_126); -lean_dec(x_124); -x_214 = lean_ctor_get(x_1, 3); -lean_inc(x_214); -if (lean_obj_tag(x_214) == 0) -{ -lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; -x_215 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4; -x_216 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_215); -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_127 = x_217; -x_128 = x_218; -goto block_213; +lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; size_t 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_267; +x_173 = lean_ctor_get(x_172, 0); +lean_inc(x_173); +x_174 = lean_ctor_get(x_172, 1); +lean_inc(x_174); +lean_dec(x_172); +x_175 = lean_ctor_get(x_157, 4); +lean_inc(x_175); +x_176 = lean_ctor_get(x_157, 3); +lean_inc(x_176); +x_177 = l_Array_zip___rarg(x_175, x_176); +lean_dec(x_176); +lean_dec(x_175); +x_178 = lean_array_get_size(x_177); +x_179 = lean_usize_of_nat(x_178); +lean_dec(x_178); +lean_inc(x_6); +x_180 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19(x_179, x_16, x_177, x_2, x_3, x_4, x_5, x_6, x_7, x_174); +lean_dec(x_3); +lean_dec(x_2); +x_181 = lean_ctor_get(x_180, 0); +lean_inc(x_181); +x_182 = lean_ctor_get(x_180, 1); +lean_inc(x_182); +if (lean_is_exclusive(x_180)) { + lean_ctor_release(x_180, 0); + lean_ctor_release(x_180, 1); + x_183 = x_180; +} else { + lean_dec_ref(x_180); + x_183 = lean_box(0); } -else -{ -lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_219 = lean_ctor_get(x_214, 0); -lean_inc(x_219); -lean_dec(x_214); -x_220 = lean_ctor_get(x_219, 0); -lean_inc(x_220); -x_221 = lean_ctor_get(x_219, 1); -lean_inc(x_221); -lean_dec(x_219); -x_127 = x_220; -x_128 = x_221; -goto block_213; +x_267 = lean_ctor_get(x_157, 2); +lean_inc(x_267); +if (lean_obj_tag(x_267) == 0) +{ +lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; +x_268 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10; +x_269 = l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23(x_268); +x_270 = lean_ctor_get(x_269, 0); +lean_inc(x_270); +x_271 = lean_ctor_get(x_269, 1); +lean_inc(x_271); +lean_dec(x_269); +x_184 = x_270; +x_185 = x_271; +goto block_266; +} +else +{ +lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_272 = lean_ctor_get(x_267, 0); +lean_inc(x_272); +lean_dec(x_267); +x_273 = lean_ctor_get(x_272, 0); +lean_inc(x_273); +x_274 = lean_ctor_get(x_272, 1); +lean_inc(x_274); +lean_dec(x_272); +x_184 = x_273; +x_185 = x_274; +goto block_266; } -block_213: +block_266: { -lean_object* x_129; uint8_t x_130; uint8_t x_131; lean_object* x_132; -x_129 = lean_ctor_get(x_7, 2); -lean_inc(x_129); -x_130 = lean_ctor_get_uint8(x_1, sizeof(void*)*8); -if (x_130 == 0) +lean_object* x_186; uint8_t x_187; uint8_t x_188; lean_object* x_189; +x_186 = lean_ctor_get(x_6, 2); +lean_inc(x_186); +x_187 = lean_ctor_get_uint8(x_157, sizeof(void*)*8); +lean_dec(x_157); +if (x_187 == 0) { -lean_object* x_205; -lean_inc(x_8); -lean_inc(x_7); +lean_object* x_258; lean_inc(x_6); -lean_inc(x_5); -x_205 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_128, x_129, x_5, x_6, x_7, x_8, x_126); -lean_dec(x_129); -if (lean_obj_tag(x_205) == 0) +x_258 = l_Lean_PrettyPrinter_Delaborator_shouldShowMotive(x_185, x_186, x_4, x_5, x_6, x_7, x_182); +lean_dec(x_186); +if (lean_obj_tag(x_258) == 0) { -lean_object* x_206; lean_object* x_207; uint8_t x_208; -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -x_207 = lean_ctor_get(x_205, 1); -lean_inc(x_207); -lean_dec(x_205); -x_208 = lean_unbox(x_206); -lean_dec(x_206); -x_131 = x_208; -x_132 = x_207; -goto block_204; +lean_object* x_259; lean_object* x_260; uint8_t x_261; +x_259 = lean_ctor_get(x_258, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_258, 1); +lean_inc(x_260); +lean_dec(x_258); +x_261 = lean_unbox(x_259); +lean_dec(x_259); +x_188 = x_261; +x_189 = x_260; +goto block_257; } else { -uint8_t x_209; -lean_dec(x_127); -lean_dec(x_125); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); -lean_dec(x_7); +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; +lean_dec(x_184); +lean_dec(x_183); +lean_dec(x_181); +lean_dec(x_173); +lean_dec(x_159); lean_dec(x_6); +x_262 = lean_ctor_get(x_258, 0); +lean_inc(x_262); +x_263 = lean_ctor_get(x_258, 1); +lean_inc(x_263); +if (lean_is_exclusive(x_258)) { + lean_ctor_release(x_258, 0); + lean_ctor_release(x_258, 1); + x_264 = x_258; +} else { + lean_dec_ref(x_258); + x_264 = lean_box(0); +} +if (lean_is_scalar(x_264)) { + x_265 = lean_alloc_ctor(1, 2, 0); +} else { + x_265 = x_264; +} +lean_ctor_set(x_265, 0, x_262); +lean_ctor_set(x_265, 1, x_263); +return x_265; +} +} +else +{ +lean_dec(x_186); +lean_dec(x_185); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_209 = !lean_is_exclusive(x_205); -if (x_209 == 0) +x_188 = x_187; +x_189 = x_182; +goto block_257; +} +block_257: { -return x_205; +if (x_188 == 0) +{ +lean_object* x_190; uint8_t x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; size_t x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; size_t x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; +lean_dec(x_184); +x_190 = lean_ctor_get(x_6, 5); +lean_inc(x_190); +lean_dec(x_6); +x_191 = 0; +x_192 = l_Lean_SourceInfo_fromRef(x_190, x_191); +x_193 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; +lean_inc(x_192); +x_194 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_194, 0, x_192); +lean_ctor_set(x_194, 1, x_193); +x_195 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_192); +x_196 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_196, 0, x_192); +lean_ctor_set(x_196, 1, x_195); +lean_ctor_set(x_196, 2, x_19); +x_197 = lean_array_get_size(x_181); +x_198 = lean_usize_of_nat(x_197); +lean_dec(x_197); +x_199 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_198, x_16, x_181); +x_200 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_201 = l_Lean_mkSepArray(x_199, x_200); +lean_dec(x_199); +x_202 = l_Array_append___rarg(x_19, x_201); +lean_inc(x_192); +x_203 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_203, 0, x_192); +lean_ctor_set(x_203, 1, x_195); +lean_ctor_set(x_203, 2, x_202); +x_204 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; +lean_inc(x_192); +x_205 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_205, 0, x_192); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Array_zip___rarg(x_173, x_159); +lean_dec(x_159); +lean_dec(x_173); +x_207 = lean_array_get_size(x_206); +x_208 = lean_usize_of_nat(x_207); +lean_dec(x_207); +x_209 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +lean_inc(x_192); +x_210 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21(x_209, x_192, x_195, x_19, x_208, x_16, x_206); +x_211 = l_Array_append___rarg(x_19, x_210); +lean_inc(x_192); +x_212 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_212, 0, x_192); +lean_ctor_set(x_212, 1, x_195); +lean_ctor_set(x_212, 2, x_211); +x_213 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; +lean_inc(x_192); +x_214 = l_Lean_Syntax_node1(x_192, x_213, x_212); +x_215 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; +lean_inc(x_196); +x_216 = l_Lean_Syntax_node6(x_192, x_215, x_194, x_196, x_196, x_203, x_205, x_214); +if (lean_is_scalar(x_183)) { + x_217 = lean_alloc_ctor(0, 2, 0); +} else { + x_217 = x_183; +} +lean_ctor_set(x_217, 0, x_216); +lean_ctor_set(x_217, 1, x_189); +return x_217; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_210 = lean_ctor_get(x_205, 0); -x_211 = lean_ctor_get(x_205, 1); -lean_inc(x_211); -lean_inc(x_210); -lean_dec(x_205); -x_212 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -return x_212; +lean_object* x_218; uint8_t x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; size_t x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; size_t x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_218 = lean_ctor_get(x_6, 5); +lean_inc(x_218); +lean_dec(x_6); +x_219 = 0; +x_220 = l_Lean_SourceInfo_fromRef(x_218, x_219); +x_221 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1; +lean_inc(x_220); +x_222 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +x_223 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_220); +x_224 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_224, 0, x_220); +lean_ctor_set(x_224, 1, x_223); +lean_ctor_set(x_224, 2, x_19); +x_225 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; +lean_inc(x_220); +x_226 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_226, 0, x_220); +lean_ctor_set(x_226, 1, x_225); +x_227 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3; +lean_inc(x_220); +x_228 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_228, 0, x_220); +lean_ctor_set(x_228, 1, x_227); +x_229 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; +lean_inc(x_220); +x_230 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_230, 0, x_220); +lean_ctor_set(x_230, 1, x_229); +x_231 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; +lean_inc(x_220); +x_232 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_232, 0, x_220); +lean_ctor_set(x_232, 1, x_231); +x_233 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; +lean_inc(x_220); +x_234 = l_Lean_Syntax_node5(x_220, x_233, x_226, x_228, x_230, x_184, x_232); +lean_inc(x_220); +x_235 = l_Lean_Syntax_node1(x_220, x_223, x_234); +x_236 = lean_array_get_size(x_181); +x_237 = lean_usize_of_nat(x_236); +lean_dec(x_236); +x_238 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_237, x_16, x_181); +x_239 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; +x_240 = l_Lean_mkSepArray(x_238, x_239); +lean_dec(x_238); +x_241 = l_Array_append___rarg(x_19, x_240); +lean_inc(x_220); +x_242 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_242, 0, x_220); +lean_ctor_set(x_242, 1, x_223); +lean_ctor_set(x_242, 2, x_241); +x_243 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; +lean_inc(x_220); +x_244 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_244, 0, x_220); +lean_ctor_set(x_244, 1, x_243); +x_245 = l_Array_zip___rarg(x_173, x_159); +lean_dec(x_159); +lean_dec(x_173); +x_246 = lean_array_get_size(x_245); +x_247 = lean_usize_of_nat(x_246); +lean_dec(x_246); +x_248 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3; +lean_inc(x_220); +x_249 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22(x_248, x_220, x_223, x_19, x_247, x_16, x_245); +x_250 = l_Array_append___rarg(x_19, x_249); +lean_inc(x_220); +x_251 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_251, 0, x_220); +lean_ctor_set(x_251, 1, x_223); +lean_ctor_set(x_251, 2, x_250); +x_252 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; +lean_inc(x_220); +x_253 = l_Lean_Syntax_node1(x_220, x_252, x_251); +x_254 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2; +x_255 = l_Lean_Syntax_node6(x_220, x_254, x_222, x_224, x_235, x_242, x_244, x_253); +if (lean_is_scalar(x_183)) { + x_256 = lean_alloc_ctor(0, 2, 0); +} else { + x_256 = x_183; } +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_189); +return x_256; } } -else -{ -lean_dec(x_129); -lean_dec(x_128); -x_131 = x_130; -x_132 = x_126; -goto block_204; } -block_204: -{ -if (x_131 == 0) +} +else { -lean_object* x_133; uint8_t x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; size_t x_142; size_t 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; size_t 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_dec(x_127); -x_133 = lean_ctor_get(x_7, 5); -lean_inc(x_133); -x_134 = 0; -x_135 = l_Lean_SourceInfo_fromRef(x_133, x_134); -x_136 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4; -lean_inc(x_135); -x_137 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); -x_138 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_139 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_135); -x_140 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_140, 0, x_135); -lean_ctor_set(x_140, 1, x_138); -lean_ctor_set(x_140, 2, x_139); -x_141 = lean_array_get_size(x_10); -x_142 = lean_usize_of_nat(x_141); -lean_dec(x_141); -x_143 = 0; -x_144 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_142, x_143, x_10); -x_145 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; -x_146 = l_Lean_mkSepArray(x_144, x_145); -lean_dec(x_144); -x_147 = l_Array_append___rarg(x_139, x_146); -lean_inc(x_135); -x_148 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_148, 0, x_135); -lean_ctor_set(x_148, 1, x_138); -lean_ctor_set(x_148, 2, x_147); -x_149 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; -lean_inc(x_135); -x_150 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_150, 0, x_135); -lean_ctor_set(x_150, 1, x_149); -x_151 = l_Array_zip___rarg(x_125, x_11); -lean_dec(x_11); -lean_dec(x_125); -x_152 = lean_array_get_size(x_151); -x_153 = lean_usize_of_nat(x_152); -lean_dec(x_152); -x_154 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; -lean_inc(x_135); -x_155 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_154, x_135, x_138, x_139, x_153, x_143, x_151); -x_156 = l_Array_append___rarg(x_139, x_155); -lean_inc(x_135); -x_157 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_157, 0, x_135); -lean_ctor_set(x_157, 1, x_138); -lean_ctor_set(x_157, 2, x_156); -x_158 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; -lean_inc(x_135); -x_159 = l_Lean_Syntax_node1(x_135, x_158, x_157); -x_160 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; -lean_inc(x_140); -x_161 = l_Lean_Syntax_node6(x_135, x_160, x_137, x_140, x_140, x_148, x_150, x_159); -x_162 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(x_1, x_161, x_3, x_4, x_5, x_6, x_7, x_8, x_132); -lean_dec(x_8); +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +lean_dec(x_159); +lean_dec(x_157); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_162; +lean_dec(x_2); +x_275 = lean_ctor_get(x_172, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_172, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + x_277 = x_172; +} else { + lean_dec_ref(x_172); + x_277 = lean_box(0); +} +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); +} else { + x_278 = x_277; +} +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; +} } else { -lean_object* x_163; uint8_t x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; size_t x_183; size_t 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; size_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; -x_163 = lean_ctor_get(x_7, 5); -lean_inc(x_163); -x_164 = 0; -x_165 = l_Lean_SourceInfo_fromRef(x_163, x_164); -x_166 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__4; -lean_inc(x_165); -x_167 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_167, 0, x_165); -lean_ctor_set(x_167, 1, x_166); -x_168 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_169 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_165); -x_170 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_170, 0, x_165); -lean_ctor_set(x_170, 1, x_168); -lean_ctor_set(x_170, 2, x_169); -x_171 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__7; -lean_inc(x_165); -x_172 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_172, 0, x_165); -lean_ctor_set(x_172, 1, x_171); -x_173 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5; -lean_inc(x_165); -x_174 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_174, 0, x_165); -lean_ctor_set(x_174, 1, x_173); -x_175 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; -lean_inc(x_165); -x_176 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_176, 0, x_165); -lean_ctor_set(x_176, 1, x_175); -x_177 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__11; -lean_inc(x_165); -x_178 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_178, 0, x_165); -lean_ctor_set(x_178, 1, x_177); -x_179 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__9; -lean_inc(x_165); -x_180 = l_Lean_Syntax_node5(x_165, x_179, x_172, x_174, x_176, x_127, x_178); -lean_inc(x_165); -x_181 = l_Lean_Syntax_node1(x_165, x_168, x_180); -x_182 = lean_array_get_size(x_10); -x_183 = lean_usize_of_nat(x_182); -lean_dec(x_182); -x_184 = 0; -x_185 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_183, x_184, x_10); -x_186 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__12; -x_187 = l_Lean_mkSepArray(x_185, x_186); -lean_dec(x_185); -x_188 = l_Array_append___rarg(x_169, x_187); -lean_inc(x_165); -x_189 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_189, 0, x_165); -lean_ctor_set(x_189, 1, x_168); -lean_ctor_set(x_189, 2, x_188); -x_190 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__6; -lean_inc(x_165); -x_191 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_191, 0, x_165); -lean_ctor_set(x_191, 1, x_190); -x_192 = l_Array_zip___rarg(x_125, x_11); -lean_dec(x_11); -lean_dec(x_125); -x_193 = lean_array_get_size(x_192); -x_194 = lean_usize_of_nat(x_193); -lean_dec(x_193); -x_195 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__3; -lean_inc(x_165); -x_196 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(x_195, x_165, x_168, x_169, x_194, x_184, x_192); -x_197 = l_Array_append___rarg(x_169, x_196); -lean_inc(x_165); -x_198 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_198, 0, x_165); -lean_ctor_set(x_198, 1, x_168); -lean_ctor_set(x_198, 2, x_197); -x_199 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__8; -lean_inc(x_165); -x_200 = l_Lean_Syntax_node1(x_165, x_199, x_198); -x_201 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__5; -x_202 = l_Lean_Syntax_node6(x_165, x_201, x_167, x_170, x_181, x_189, x_191, x_200); -x_203 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(x_1, x_202, x_3, x_4, x_5, x_6, x_7, x_8, x_132); -lean_dec(x_8); +lean_object* x_279; uint8_t 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_dec(x_159); lean_dec(x_7); -lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_203; -} +lean_dec(x_2); +x_279 = lean_ctor_get(x_6, 5); +lean_inc(x_279); +lean_dec(x_6); +x_280 = 0; +x_281 = l_Lean_SourceInfo_fromRef(x_279, x_280); +x_282 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; +lean_inc(x_281); +x_283 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_283, 0, x_281); +lean_ctor_set(x_283, 1, x_282); +x_284 = lean_ctor_get(x_157, 3); +lean_inc(x_284); +lean_dec(x_157); +x_285 = l_Lean_PrettyPrinter_Delaborator_delabConst___closed__11; +x_286 = l_Lean_Syntax_SepArray_ofElems(x_285, x_284); +lean_dec(x_284); +x_287 = l_Array_append___rarg(x_19, x_286); +x_288 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +lean_inc(x_281); +x_289 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_289, 0, x_281); +lean_ctor_set(x_289, 1, x_288); +lean_ctor_set(x_289, 2, x_287); +x_290 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12; +x_291 = l_Lean_Syntax_node2(x_281, x_290, x_283, x_289); +x_292 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_292, 0, x_291); +lean_ctor_set(x_292, 1, x_158); +return x_292; } } } else { -uint8_t x_222; -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_8); +uint8_t x_293; 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_222 = !lean_is_exclusive(x_124); -if (x_222 == 0) +lean_dec(x_2); +x_293 = !lean_is_exclusive(x_20); +if (x_293 == 0) { -return x_124; +return x_20; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_124, 0); -x_224 = lean_ctor_get(x_124, 1); -lean_inc(x_224); -lean_inc(x_223); -lean_dec(x_124); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +lean_object* x_294; lean_object* x_295; lean_object* x_296; +x_294 = lean_ctor_get(x_20, 0); +x_295 = lean_ctor_get(x_20, 1); +lean_inc(x_295); +lean_inc(x_294); +lean_dec(x_20); +x_296 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_296, 0, x_294); +lean_ctor_set(x_296, 1, x_295); +return x_296; } } } -else +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1() { +_start: { -lean_object* x_226; lean_object* x_227; uint8_t x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_226 = lean_array_fget(x_10, x_122); -lean_dec(x_10); -x_227 = lean_ctor_get(x_7, 5); -lean_inc(x_227); -lean_dec(x_7); -x_228 = 0; -x_229 = l_Lean_SourceInfo_fromRef(x_227, x_228); -x_230 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10; -lean_inc(x_229); -x_231 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_231, 0, x_229); -lean_ctor_set(x_231, 1, x_230); -x_232 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -lean_inc(x_229); -x_233 = l_Lean_Syntax_node1(x_229, x_232, x_226); -x_234 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11; -x_235 = l_Lean_Syntax_node2(x_229, x_234, x_231, x_233); -x_236 = lean_ctor_get(x_1, 7); -lean_inc(x_236); -lean_dec(x_1); -x_237 = l_Lean_Syntax_mkApp(x_235, x_236); -x_238 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_238, 0, x_237); -lean_ctor_set(x_238, 1, x_9); -return x_238; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3), 8, 0); +return x_1; } } +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6), 8, 0); +return x_1; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { _start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_1, 4); -lean_inc(x_9); -x_10 = lean_array_get_size(x_9); -lean_dec(x_9); -x_11 = lean_ctor_get(x_1, 0); +lean_object* x_9; +x_9 = l_Lean_Expr_getAppFn(x_1); +if (lean_obj_tag(x_9) == 4) +{ +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); +x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -x_13 = lean_nat_dec_lt(x_10, x_12); -lean_dec(x_12); -lean_dec(x_10); -if (x_13 == 0) +lean_dec(x_9); +lean_inc(x_10); +x_12 = l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__1(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +if (lean_obj_tag(x_13) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_14 = lean_ctor_get(x_1, 6); -lean_inc(x_14); -x_15 = lean_array_get_size(x_14); -lean_dec(x_14); -x_16 = lean_ctor_get(x_11, 2); -lean_inc(x_16); +lean_object* x_14; lean_object* x_15; lean_dec(x_11); -x_17 = lean_array_get_size(x_16); -lean_dec(x_16); -x_18 = lean_nat_dec_lt(x_15, x_17); -lean_dec(x_17); -lean_dec(x_15); -if (x_18 == 0) -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_box(0); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6(x_1, x_19, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -return x_20; -} -else -{ -lean_object* x_21; uint8_t x_22; +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_21 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -x_22 = !lean_is_exclusive(x_21); -if (x_22 == 0) -{ -return x_21; +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_14); +return x_15; } else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_21, 0); -x_24 = lean_ctor_get(x_21, 1); -lean_inc(x_24); -lean_inc(x_23); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_ctor_get(x_13, 0); +lean_inc(x_17); +lean_dec(x_13); +x_18 = l_Lean_Meta_Match_MatcherInfo_arity(x_17); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +x_20 = lean_unsigned_to_nat(1u); +x_21 = lean_nat_add(x_20, x_19); +lean_dec(x_19); +x_22 = l_Lean_Meta_Match_MatcherInfo_numAlts(x_17); +x_23 = lean_nat_add(x_21, x_22); +lean_dec(x_22); lean_dec(x_21); -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; -} +x_24 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1___boxed), 11, 3); +lean_closure_set(x_24, 0, x_10); +lean_closure_set(x_24, 1, x_11); +lean_closure_set(x_24, 2, x_17); +x_25 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__1; +x_26 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); +lean_closure_set(x_26, 0, x_25); +lean_closure_set(x_26, 1, x_24); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1; +x_28 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBoundedAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__6), 10, 3); +lean_closure_set(x_28, 0, x_23); +lean_closure_set(x_28, 1, x_26); +lean_closure_set(x_28, 2, x_27); +x_29 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2; +x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); +lean_closure_set(x_30, 0, x_28); +lean_closure_set(x_30, 1, x_29); +x_31 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_18, x_30, x_2, x_3, x_4, x_5, x_6, x_7, x_16); +lean_dec(x_18); +return x_31; } } else { -lean_object* x_26; uint8_t x_27; -lean_dec(x_11); +lean_object* x_32; +lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -lean_dec(x_1); -x_26 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) -{ -return x_26; -} -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} +x_32 = l_Lean_PrettyPrinter_Delaborator_failure___rarg(x_8); +return x_32; } } } @@ -25702,7 +25729,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___boxed), 8, 0); return x_1; } } @@ -25722,46 +25749,30 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___close _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_getPPMatch___boxed), 1, 0); return x_1; } } static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4() { _start: { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7), 8, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_getPPMatch___boxed), 1, 0); -return x_1; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2; +x_3 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_whenPPOption), 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_delabAppMatch(lean_object* x_1, 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; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2; -x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3; -x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFnArgs___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__5), 9, 2); -lean_closure_set(x_10, 0, x_8); -lean_closure_set(x_10, 1, x_9); -x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4; -x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); -lean_closure_set(x_12, 0, x_10); -lean_closure_set(x_12, 1, x_11); -x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5; -x_14 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_whenPPOption), 9, 2); -lean_closure_set(x_14, 0, x_13); -lean_closure_set(x_14, 1, x_12); -x_15 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; -x_16 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_15, x_14, x_1, x_2, x_3, x_4, x_5, x_6, x_7); -return x_16; +lean_object* x_8; lean_object* x_9; lean_object* x_10; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; +x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4; +x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); +return x_10; } } LEAN_EXPORT lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { @@ -25816,7 +25827,7 @@ x_12 = l_Lean_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_delabAppM return x_12; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { size_t x_4; size_t x_5; lean_object* x_6; @@ -25824,37 +25835,85 @@ x_4 = lean_unbox_usize(x_1); lean_dec(x_1); x_5 = lean_unbox_usize(x_2); lean_dec(x_2); -x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__9(x_4, x_5, x_3); +x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); -lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); -lean_dec(x_6); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +size_t x_5; size_t x_6; lean_object* x_7; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12(x_1, x_5, x_6, x_4); lean_dec(x_1); -return x_10; +return x_7; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -size_t x_8; size_t x_9; lean_object* x_10; -x_8 = lean_unbox_usize(x_5); +lean_object* x_4; +x_4 = l_Array_filterMapM___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_1, x_2, x_3); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, 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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +lean_dec(x_7); lean_dec(x_5); -x_9 = lean_unbox_usize(x_6); +lean_dec(x_4); +return x_14; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: +{ +lean_object* x_16; +x_16 = l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_16; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +size_t x_11; size_t x_12; lean_object* x_13; +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_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19(x_11, x_12, 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); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__11(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +lean_dec(x_5); +lean_dec(x_4); +return x_13; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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; +x_4 = lean_unbox_usize(x_1); lean_dec(x_1); -return x_10; +x_5 = lean_unbox_usize(x_2); +lean_dec(x_2); +x_6 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__20(x_4, x_5, x_3); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25862,12 +25921,12 @@ x_8 = lean_unbox_usize(x_5); lean_dec(x_5); x_9 = lean_unbox_usize(x_6); lean_dec(x_6); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__13(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -25875,21 +25934,19 @@ x_8 = lean_unbox_usize(x_5); lean_dec(x_5); x_9 = lean_unbox_usize(x_6); lean_dec(x_6); -x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__14(x_1, x_2, x_3, x_4, x_8, x_9, x_7); +x_10 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__22(x_1, x_2, x_3, x_4, x_8, x_9, x_7); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_object* x_12; +x_12 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; +return x_12; } } LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { @@ -25906,18 +25963,23 @@ lean_dec(x_3); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__5(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_object* x_11; +x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_4); -lean_dec(x_3); -return x_10; +lean_dec(x_1); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabAppMatch___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) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_1); +return x_9; } } static lean_object* _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1() { @@ -25960,24 +26022,7 @@ x_6 = l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(x_2, x_3, x_4, x_5, x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_9 = l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_getExprKind___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8); -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_ctor_get(x_9, 1); -lean_inc(x_11); -lean_dec(x_9); -x_12 = l_Lean_Expr_appArg_x21(x_10); -lean_dec(x_10); -x_13 = lean_unsigned_to_nat(1u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delab___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); -return x_14; -} -} -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -25993,7 +26038,7 @@ x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_ return x_14; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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; @@ -26049,7 +26094,7 @@ return x_27; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___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_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2___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_object* x_11) { _start: { lean_object* x_12; @@ -26076,7 +26121,7 @@ lean_dec(x_4); lean_dec(x_15); x_17 = lean_apply_1(x_3, x_13); x_18 = lean_unsigned_to_nat(1u); -x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__4(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(x_16, x_18, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_14); return x_19; } else @@ -26112,7 +26157,7 @@ return x_23; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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) { _start: { 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; @@ -26124,7 +26169,7 @@ lean_inc(x_13); lean_dec(x_11); x_14 = l_Lean_Expr_binderInfo(x_12); x_15 = l_Lean_Expr_bindingDomain_x21(x_12); -x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___rarg___lambda__1), 11, 3); +x_16 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2___rarg___lambda__1), 11, 3); lean_closure_set(x_16, 0, x_2); lean_closure_set(x_16, 1, x_12); lean_closure_set(x_16, 2, x_3); @@ -26133,15 +26178,15 @@ x_18 = l_Lean_Meta_withLocalDecl___at_Lean_PrettyPrinter_Delaborator_withBinding return x_18; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___rarg), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2___rarg), 10, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; @@ -26154,11 +26199,11 @@ lean_dec(x_9); x_12 = l_Lean_Expr_appArg_x21(x_10); lean_dec(x_10); x_13 = lean_unsigned_to_nat(1u); -x_14 = l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__4(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_delabLetFun___spec__3(x_12, x_13, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_11); return x_14; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; @@ -26275,7 +26320,86 @@ return x_49; } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_PrettyPrinter_Delaborator_delab(x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_1); +lean_ctor_set(x_12, 1, x_11); +lean_ctor_set(x_9, 0, x_12); +return x_9; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_13 = lean_ctor_get(x_9, 0); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +lean_inc(x_13); +lean_dec(x_9); +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_1); +lean_ctor_set(x_15, 1, x_13); +x_16 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +return x_16; +} +} +else +{ +uint8_t x_17; +lean_dec(x_1); +x_17 = !lean_is_exclusive(x_9); +if (x_17 == 0) +{ +return x_9; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_9, 0); +x_19 = lean_ctor_get(x_9, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_9); +x_20 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1), 8, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1), 8, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3() { _start: { lean_object* x_1; @@ -26283,19 +26407,19 @@ x_1 = lean_mk_string_from_bytes("let_fun", 7); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___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_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_delabLetFun___lambda__1___closed__1; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___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_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__5() { _start: { lean_object* x_1; @@ -26303,19 +26427,19 @@ x_1 = lean_mk_string_from_bytes("letDecl", 7); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_PrettyPrinter_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_delabLetFun___lambda__1___closed__3; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___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_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7() { _start: { lean_object* x_1; @@ -26323,19 +26447,19 @@ x_1 = lean_mk_string_from_bytes("letIdDecl", 9); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_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_delabLetFun___lambda__1___closed__5; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9() { _start: { lean_object* x_1; @@ -26343,7 +26467,7 @@ x_1 = lean_mk_string_from_bytes(";", 1); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__8() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10() { _start: { lean_object* x_1; @@ -26351,19 +26475,19 @@ x_1 = lean_mk_string_from_bytes("typeSpec", 8); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___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_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_delabLetFun___lambda__1___closed__8; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12() { _start: { lean_object* x_1; @@ -26371,7 +26495,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_getPPLetVarTypes___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13() { _start: { lean_object* x_1; @@ -26379,7 +26503,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_getPPAnalysisLetVarType___boxed), 1, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___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; uint8_t x_12; lean_object* x_13; @@ -26443,73 +26567,71 @@ lean_inc(x_2); x_17 = l_Lean_PrettyPrinter_Delaborator_getUnusedName(x_15, x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_13); if (lean_obj_tag(x_17) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; x_18 = lean_ctor_get(x_17, 0); lean_inc(x_18); x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); lean_dec(x_17); -x_20 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_21 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); -lean_closure_set(x_21, 0, x_20); +x_20 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_22 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_19); -if (lean_obj_tag(x_22) == 0) +x_21 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_19); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -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; +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); +lean_dec(x_21); lean_inc(x_18); -x_25 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent___boxed), 9, 1); -lean_closure_set(x_25, 0, x_18); -x_26 = l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2; -x_27 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__3___rarg), 10, 3); -lean_closure_set(x_27, 0, x_18); -lean_closure_set(x_27, 1, x_25); -lean_closure_set(x_27, 2, x_26); +x_24 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_mkAnnotatedIdent___boxed), 9, 1); +lean_closure_set(x_24, 0, x_18); +x_25 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2; +x_26 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2___rarg), 10, 3); +lean_closure_set(x_26, 0, x_18); +lean_closure_set(x_26, 1, x_24); +lean_closure_set(x_26, 2, x_25); 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_28 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_24); -if (lean_obj_tag(x_28) == 0) +x_27 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__4(x_26, x_2, x_3, x_4, x_5, x_6, x_7, x_23); +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; uint8_t x_34; lean_object* x_35; lean_object* x_112; lean_object* x_113; -x_29 = lean_ctor_get(x_28, 0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; lean_object* x_112; lean_object* x_113; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); lean_inc(x_29); -x_30 = lean_ctor_get(x_28, 1); -lean_inc(x_30); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_31 = x_28; +if (lean_is_exclusive(x_27)) { + lean_ctor_release(x_27, 0); + lean_ctor_release(x_27, 1); + x_30 = x_27; } else { - lean_dec_ref(x_28); - x_31 = lean_box(0); + lean_dec_ref(x_27); + x_30 = lean_box(0); } -x_32 = lean_ctor_get(x_29, 0); +x_31 = lean_ctor_get(x_28, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_28, 1); lean_inc(x_32); -x_33 = lean_ctor_get(x_29, 1); -lean_inc(x_33); -lean_dec(x_29); -x_112 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10; +lean_dec(x_28); +x_112 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12; 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_113 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_30); +x_113 = l_Lean_PrettyPrinter_Delaborator_getPPOption___rarg(x_112, x_2, x_3, x_4, x_5, x_6, x_7, x_29); if (lean_obj_tag(x_113) == 0) { lean_object* x_114; uint8_t x_115; @@ -26523,7 +26645,7 @@ lean_dec(x_114); x_116 = lean_ctor_get(x_113, 1); lean_inc(x_116); lean_dec(x_113); -x_117 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11; +x_117 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); @@ -26541,17 +26663,17 @@ lean_inc(x_120); lean_dec(x_118); x_121 = lean_unbox(x_119); lean_dec(x_119); -x_34 = x_121; -x_35 = x_120; +x_33 = x_121; +x_34 = x_120; goto block_111; } else { uint8_t x_122; -lean_dec(x_33); lean_dec(x_32); lean_dec(x_31); -lean_dec(x_23); +lean_dec(x_30); +lean_dec(x_22); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -26586,18 +26708,18 @@ lean_inc(x_126); lean_dec(x_113); x_127 = lean_unbox(x_114); lean_dec(x_114); -x_34 = x_127; -x_35 = x_126; +x_33 = x_127; +x_34 = x_126; goto block_111; } } else { uint8_t x_128; -lean_dec(x_33); lean_dec(x_32); lean_dec(x_31); -lean_dec(x_23); +lean_dec(x_30); +lean_dec(x_22); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -26625,65 +26747,66 @@ return x_131; } block_111: { -if (x_34 == 0) +if (x_33 == 0) { -lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; +lean_object* x_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; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_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_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_36 = lean_ctor_get(x_6, 5); -lean_inc(x_36); +x_35 = lean_ctor_get(x_6, 5); +lean_inc(x_35); lean_dec(x_6); -x_37 = 0; -x_38 = l_Lean_SourceInfo_fromRef(x_36, x_37); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1; -lean_inc(x_38); -x_40 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; -x_42 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; -lean_inc(x_38); -x_43 = lean_alloc_ctor(1, 3, 0); -lean_ctor_set(x_43, 0, x_38); -lean_ctor_set(x_43, 1, x_41); -lean_ctor_set(x_43, 2, x_42); -x_44 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; -lean_inc(x_38); -x_45 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_45, 0, x_38); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; -lean_inc(x_43); -lean_inc(x_38); -x_47 = l_Lean_Syntax_node5(x_38, x_46, x_32, x_43, x_43, x_45, x_23); -x_48 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; -lean_inc(x_38); -x_49 = l_Lean_Syntax_node1(x_38, x_48, x_47); -x_50 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; -lean_inc(x_38); -x_51 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_51, 0, x_38); -lean_ctor_set(x_51, 1, x_50); -x_52 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2; -x_53 = l_Lean_Syntax_node4(x_38, x_52, x_40, x_49, x_51, x_33); -if (lean_is_scalar(x_31)) { - x_54 = lean_alloc_ctor(0, 2, 0); +x_36 = 0; +x_37 = l_Lean_SourceInfo_fromRef(x_35, x_36); +x_38 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3; +lean_inc(x_37); +x_39 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_39, 0, x_37); +lean_ctor_set(x_39, 1, x_38); +x_40 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__10; +x_41 = l_Lean_PrettyPrinter_Delaborator_delabSort___closed__9; +lean_inc(x_37); +x_42 = lean_alloc_ctor(1, 3, 0); +lean_ctor_set(x_42, 0, x_37); +lean_ctor_set(x_42, 1, x_40); +lean_ctor_set(x_42, 2, x_41); +x_43 = l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_unexpandStructureInstance___spec__1___lambda__1___closed__5; +lean_inc(x_37); +x_44 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_44, 0, x_37); +lean_ctor_set(x_44, 1, x_43); +x_45 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; +lean_inc(x_42); +lean_inc(x_37); +x_46 = l_Lean_Syntax_node5(x_37, x_45, x_31, x_42, x_42, x_44, x_22); +x_47 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; +lean_inc(x_37); +x_48 = l_Lean_Syntax_node1(x_37, x_47, x_46); +x_49 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; +lean_inc(x_37); +x_50 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_50, 0, x_37); +lean_ctor_set(x_50, 1, x_49); +x_51 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4; +x_52 = l_Lean_Syntax_node4(x_37, x_51, x_39, x_48, x_50, x_32); +if (lean_is_scalar(x_30)) { + x_53 = lean_alloc_ctor(0, 2, 0); } else { - x_54 = x_31; + x_53 = x_30; } -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_35); -return x_54; +lean_ctor_set(x_53, 0, x_52); +lean_ctor_set(x_53, 1, x_34); +return x_53; } else { -lean_object* x_55; -lean_dec(x_31); +lean_object* x_54; lean_object* x_55; +lean_dec(x_30); +x_54 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; lean_inc(x_6); -x_55 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_9, x_20, x_2, x_3, x_4, x_5, x_6, x_7, x_35); +x_55 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_9, x_54, x_2, x_3, x_4, x_5, x_6, x_7, x_34); if (lean_obj_tag(x_55) == 0) { uint8_t x_56; @@ -26697,7 +26820,7 @@ lean_inc(x_58); lean_dec(x_6); x_59 = 0; x_60 = l_Lean_SourceInfo_fromRef(x_58, x_59); -x_61 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1; +x_61 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3; lean_inc(x_60); x_62 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_62, 0, x_60); @@ -26714,7 +26837,7 @@ lean_inc(x_60); x_67 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_67, 0, x_60); lean_ctor_set(x_67, 1, x_66); -x_68 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_68 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_60); x_69 = l_Lean_Syntax_node2(x_60, x_68, x_67, x_57); lean_inc(x_60); @@ -26724,19 +26847,19 @@ lean_inc(x_60); x_72 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_72, 0, x_60); lean_ctor_set(x_72, 1, x_71); -x_73 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_73 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_60); -x_74 = l_Lean_Syntax_node5(x_60, x_73, x_32, x_65, x_70, x_72, x_23); -x_75 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_74 = l_Lean_Syntax_node5(x_60, x_73, x_31, x_65, x_70, x_72, x_22); +x_75 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_60); x_76 = l_Lean_Syntax_node1(x_60, x_75, x_74); -x_77 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; +x_77 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; lean_inc(x_60); x_78 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_78, 0, x_60); lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2; -x_80 = l_Lean_Syntax_node4(x_60, x_79, x_62, x_76, x_78, x_33); +x_79 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4; +x_80 = l_Lean_Syntax_node4(x_60, x_79, x_62, x_76, x_78, x_32); lean_ctor_set(x_55, 0, x_80); return x_55; } @@ -26753,7 +26876,7 @@ lean_inc(x_83); lean_dec(x_6); x_84 = 0; x_85 = l_Lean_SourceInfo_fromRef(x_83, x_84); -x_86 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1; +x_86 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3; lean_inc(x_85); x_87 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_87, 0, x_85); @@ -26770,7 +26893,7 @@ lean_inc(x_85); x_92 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_92, 0, x_85); lean_ctor_set(x_92, 1, x_91); -x_93 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_93 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_85); x_94 = l_Lean_Syntax_node2(x_85, x_93, x_92, x_81); lean_inc(x_85); @@ -26780,19 +26903,19 @@ lean_inc(x_85); x_97 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_97, 0, x_85); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_98 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_85); -x_99 = l_Lean_Syntax_node5(x_85, x_98, x_32, x_90, x_95, x_97, x_23); -x_100 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_99 = l_Lean_Syntax_node5(x_85, x_98, x_31, x_90, x_95, x_97, x_22); +x_100 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_85); x_101 = l_Lean_Syntax_node1(x_85, x_100, x_99); -x_102 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; +x_102 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; lean_inc(x_85); x_103 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_103, 0, x_85); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2; -x_105 = l_Lean_Syntax_node4(x_85, x_104, x_87, x_101, x_103, x_33); +x_104 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4; +x_105 = l_Lean_Syntax_node4(x_85, x_104, x_87, x_101, x_103, x_32); x_106 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_106, 0, x_105); lean_ctor_set(x_106, 1, x_82); @@ -26802,9 +26925,9 @@ return x_106; else { uint8_t x_107; -lean_dec(x_33); lean_dec(x_32); -lean_dec(x_23); +lean_dec(x_31); +lean_dec(x_22); lean_dec(x_6); x_107 = !lean_is_exclusive(x_55); if (x_107 == 0) @@ -26831,26 +26954,26 @@ return x_110; else { uint8_t x_132; -lean_dec(x_23); +lean_dec(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_2); -x_132 = !lean_is_exclusive(x_28); +x_132 = !lean_is_exclusive(x_27); if (x_132 == 0) { -return x_28; +return x_27; } else { lean_object* x_133; lean_object* x_134; lean_object* x_135; -x_133 = lean_ctor_get(x_28, 0); -x_134 = lean_ctor_get(x_28, 1); +x_133 = lean_ctor_get(x_27, 0); +x_134 = lean_ctor_get(x_27, 1); lean_inc(x_134); lean_inc(x_133); -lean_dec(x_28); +lean_dec(x_27); x_135 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_135, 0, x_133); lean_ctor_set(x_135, 1, x_134); @@ -26868,19 +26991,19 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_136 = !lean_is_exclusive(x_22); +x_136 = !lean_is_exclusive(x_21); if (x_136 == 0) { -return x_22; +return x_21; } else { lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_22, 0); -x_138 = lean_ctor_get(x_22, 1); +x_137 = lean_ctor_get(x_21, 0); +x_138 = lean_ctor_get(x_21, 1); lean_inc(x_138); lean_inc(x_137); -lean_dec(x_22); +lean_dec(x_21); x_139 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); @@ -26937,7 +27060,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed_ _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___boxed), 8, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___boxed), 8, 0); return x_1; } } @@ -26969,26 +27092,26 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun(lean_objec _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_9 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__3; x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___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) { _start: { lean_object* x_10; -x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_1); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; -x_9 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); lean_dec(x_1); return x_9; } @@ -27357,8 +27480,8 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__ _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); +x_1 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -27366,34 +27489,24 @@ return x_2; static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabMData___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; -} -} -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3() { -_start: -{ lean_object* x_1; x_1 = lean_mk_string_from_bytes("inaccessible", 12); return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___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_delabMData___closed__3; +x_4 = l_Lean_PrettyPrinter_Delaborator_delabMData___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_delabMData___closed__5() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4() { _start: { lean_object* x_1; @@ -27437,7 +27550,7 @@ else lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_dec(x_15); x_19 = lean_box(0); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; +x_20 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; x_21 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withMDataExpr___at_Lean_PrettyPrinter_Delaborator_delabMData___spec__1(x_19, x_20, x_1, x_2, x_3, x_4, x_5, x_6, x_14); return x_21; } @@ -27492,7 +27605,7 @@ lean_inc(x_32); lean_dec(x_5); x_33 = 0; x_34 = l_Lean_SourceInfo_fromRef(x_32, x_33); -x_35 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5; +x_35 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; lean_inc(x_34); x_36 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_36, 0, x_34); @@ -27502,7 +27615,7 @@ lean_inc(x_34); x_38 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_38, 0, x_34); lean_ctor_set(x_38, 1, x_37); -x_39 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; +x_39 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3; x_40 = l_Lean_Syntax_node3(x_34, x_39, x_36, x_31, x_38); lean_ctor_set(x_24, 0, x_40); return x_24; @@ -27520,7 +27633,7 @@ lean_inc(x_43); lean_dec(x_5); x_44 = 0; x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); -x_46 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5; +x_46 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; lean_inc(x_45); x_47 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_47, 0, x_45); @@ -27530,7 +27643,7 @@ lean_inc(x_45); x_49 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_49, 0, x_45); lean_ctor_set(x_49, 1, x_48); -x_50 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4; +x_50 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3; x_51 = l_Lean_Syntax_node3(x_45, x_50, x_47, x_41, x_49); x_52 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_52, 0, x_51); @@ -27858,9 +27971,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__2___boxed), 9, 1); +x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); lean_closure_set(x_10, 0, x_2); -x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1; +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_shouldGroupWithNext___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } @@ -29241,7 +29354,7 @@ x_32 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_32, 0, x_20); lean_ctor_set(x_32, 1, x_30); lean_ctor_set(x_32, 2, x_28); -x_33 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; +x_33 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; lean_inc(x_20); x_34 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_34, 0, x_20); @@ -29289,7 +29402,7 @@ x_54 = lean_alloc_ctor(1, 3, 0); lean_ctor_set(x_54, 0, x_42); lean_ctor_set(x_54, 1, x_52); lean_ctor_set(x_54, 2, x_50); -x_55 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4; +x_55 = l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4; lean_inc(x_42); x_56 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_56, 0, x_42); @@ -32321,7 +32434,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLetE___closed__1; -x_3 = lean_unsigned_to_nat(854u); +x_3 = lean_unsigned_to_nat(927u); x_4 = lean_unsigned_to_nat(38u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -32437,7 +32550,7 @@ lean_inc(x_29); x_30 = lean_ctor_get(x_26, 1); lean_inc(x_30); lean_dec(x_26); -x_110 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10; +x_110 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); @@ -32458,7 +32571,7 @@ lean_dec(x_112); x_114 = lean_ctor_get(x_111, 1); lean_inc(x_114); lean_dec(x_111); -x_115 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11; +x_115 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13; lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); @@ -32593,14 +32706,14 @@ lean_inc(x_35); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_35); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_43 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_40); lean_inc(x_35); x_44 = l_Lean_Syntax_node5(x_35, x_43, x_29, x_40, x_40, x_42, x_21); -x_45 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_45 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_35); x_46 = l_Lean_Syntax_node1(x_35, x_45, x_44); -x_47 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; +x_47 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; lean_inc(x_35); x_48 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_48, 0, x_35); @@ -32653,7 +32766,7 @@ lean_inc(x_58); x_65 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_65, 0, x_58); lean_ctor_set(x_65, 1, x_64); -x_66 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_66 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_58); x_67 = l_Lean_Syntax_node2(x_58, x_66, x_65, x_55); lean_inc(x_58); @@ -32663,13 +32776,13 @@ lean_inc(x_58); x_70 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_70, 0, x_58); lean_ctor_set(x_70, 1, x_69); -x_71 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_71 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_58); x_72 = l_Lean_Syntax_node5(x_58, x_71, x_29, x_63, x_68, x_70, x_21); -x_73 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_73 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_58); x_74 = l_Lean_Syntax_node1(x_58, x_73, x_72); -x_75 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; +x_75 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; lean_inc(x_58); x_76 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_76, 0, x_58); @@ -32709,7 +32822,7 @@ lean_inc(x_83); x_90 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_90, 0, x_83); lean_ctor_set(x_90, 1, x_89); -x_91 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_91 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_83); x_92 = l_Lean_Syntax_node2(x_83, x_91, x_90, x_79); lean_inc(x_83); @@ -32719,13 +32832,13 @@ lean_inc(x_83); x_95 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_95, 0, x_83); lean_ctor_set(x_95, 1, x_94); -x_96 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_96 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_83); x_97 = l_Lean_Syntax_node5(x_83, x_96, x_29, x_88, x_93, x_95, x_21); -x_98 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_98 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_83); x_99 = l_Lean_Syntax_node1(x_83, x_98, x_97); -x_100 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7; +x_100 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9; lean_inc(x_83); x_101 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_101, 0, x_83); @@ -33144,7 +33257,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabLit___closed__2( lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabLit___closed__1; -x_3 = lean_unsigned_to_nat(874u); +x_3 = lean_unsigned_to_nat(947u); x_4 = lean_unsigned_to_nat(29u); 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); @@ -33590,7 +33703,7 @@ lean_dec(x_32); x_39 = lean_unsigned_to_nat(0u); x_40 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; lean_inc(x_6); -x_41 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_39, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_38); +x_41 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_39, x_40, x_2, x_3, x_4, x_5, x_6, x_7, x_38); if (lean_obj_tag(x_41) == 0) { uint8_t x_42; @@ -33981,7 +34094,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_17 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_15); +x_17 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_16, x_2, x_3, x_4, x_5, x_6, x_7, x_15); if (lean_obj_tag(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; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -34042,7 +34155,7 @@ lean_inc(x_33); lean_dec(x_27); x_34 = lean_unsigned_to_nat(0u); x_35 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_36 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_34, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_33); +x_36 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_34, x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_33); if (lean_obj_tag(x_36) == 0) { uint8_t x_37; @@ -34235,7 +34348,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___clo { lean_object* x_1; lean_object* x_2; x_1 = l_Lean_PrettyPrinter_Delaborator_delabDivRatCore___closed__5; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -34424,7 +34537,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_34 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_32); +x_34 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_33, x_2, x_3, x_4, x_5, x_6, x_7, x_32); if (lean_obj_tag(x_34) == 0) { lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; @@ -34440,7 +34553,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_38 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_36); +x_38 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_37, x_2, x_3, x_4, x_5, x_6, x_7, x_36); if (lean_obj_tag(x_38) == 0) { 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; @@ -34500,7 +34613,7 @@ x_54 = lean_ctor_get(x_48, 1); lean_inc(x_54); lean_dec(x_48); x_55 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_56 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_16, x_55, x_2, x_3, x_4, x_5, x_6, x_7, x_54); +x_56 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_16, x_55, x_2, x_3, x_4, x_5, x_6, x_7, x_54); if (lean_obj_tag(x_56) == 0) { uint8_t x_57; @@ -35073,7 +35186,7 @@ x_36 = lean_string_utf8_extract(x_11, x_33, x_22); lean_dec(x_22); lean_dec(x_33); lean_dec(x_11); -x_37 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +x_37 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; x_38 = lean_string_append(x_35, x_37); x_39 = lean_string_append(x_38, x_36); lean_dec(x_36); @@ -35958,7 +36071,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabProj___closed__2 lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabProj___closed__1; -x_3 = lean_unsigned_to_nat(980u); +x_3 = lean_unsigned_to_nat(1053u); x_4 = lean_unsigned_to_nat(36u); 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); @@ -36023,12 +36136,12 @@ lean_inc(x_23); lean_dec(x_5); x_24 = 0; x_25 = l_Lean_SourceInfo_fromRef(x_23, x_24); -x_26 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +x_26 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; 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__6___closed__2; +x_28 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; x_29 = l_Lean_Syntax_node3(x_25, x_28, x_16, x_27, x_22); lean_ctor_set(x_14, 0, x_29); return x_14; @@ -36053,12 +36166,12 @@ lean_inc(x_38); lean_dec(x_5); x_39 = 0; x_40 = l_Lean_SourceInfo_fromRef(x_38, x_39); -x_41 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3; +x_41 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; lean_inc(x_40); x_42 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_42, 0, x_40); lean_ctor_set(x_42, 1, x_41); -x_43 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2; +x_43 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2; x_44 = l_Lean_Syntax_node3(x_40, x_43, x_30, x_42, x_37); x_45 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_45, 0, x_44); @@ -36109,7 +36222,7 @@ static lean_object* _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabPro { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1; +x_2 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -36372,7 +36485,7 @@ x_19 = lean_ctor_get(x_15, 1); lean_inc(x_19); lean_dec(x_15); x_20 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_21 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6___boxed), 9, 2); +x_21 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); lean_closure_set(x_21, 0, x_19); lean_closure_set(x_21, 1, x_20); x_22 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__1; @@ -36402,7 +36515,7 @@ x_30 = lean_ctor_get(x_15, 1); lean_inc(x_30); lean_dec(x_15); x_31 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_32 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6___boxed), 9, 2); +x_32 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); lean_closure_set(x_32, 0, x_30); lean_closure_set(x_32, 1, x_31); x_33 = l_Lean_PrettyPrinter_Delaborator_withOverApp(x_26, x_32, x_2, x_3, x_4, x_5, x_6, x_7, x_25); @@ -36416,7 +36529,7 @@ x_34 = lean_ctor_get(x_15, 1); lean_inc(x_34); lean_dec(x_15); x_35 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_36 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6___boxed), 9, 2); +x_36 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); lean_closure_set(x_36, 0, x_34); lean_closure_set(x_36, 1, x_35); x_37 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__2; @@ -36440,7 +36553,7 @@ x_42 = lean_ctor_get(x_15, 1); lean_inc(x_42); lean_dec(x_15); x_43 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; -x_44 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6___boxed), 9, 2); +x_44 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5___boxed), 9, 2); lean_closure_set(x_44, 0, x_42); lean_closure_set(x_44, 1, x_43); x_45 = l_Lean_PrettyPrinter_Delaborator_coeDelaborator___lambda__4___closed__3; @@ -36681,9 +36794,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBod _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___lambda__2___boxed), 9, 1); +x_10 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___lambda__2___boxed), 9, 1); lean_closure_set(x_10, 0, x_2); -x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1; +x_11 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1; x_12 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_delabDIte_delabBranch___spec__2___rarg(x_1, x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_12; } @@ -36992,8 +37105,8 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDIte___lambda__1 _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; -x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2), 8, 1); +x_1 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -37121,7 +37234,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); +x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; @@ -37377,7 +37490,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDIte(lean_object* _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_9 = l_Lean_PrettyPrinter_Delaborator_delabDIte___closed__3; x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; @@ -37522,14 +37635,14 @@ goto block_63; block_63: { lean_object* x_14; lean_object* x_15; -x_14 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; +x_14 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); +x_15 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_14, x_2, x_3, x_4, x_5, x_6, x_7, x_13); if (lean_obj_tag(x_15) == 0) { lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; @@ -37538,14 +37651,14 @@ lean_inc(x_16); x_17 = lean_ctor_get(x_15, 1); lean_inc(x_17); lean_dec(x_15); -x_18 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; +x_18 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_18, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_18, x_2, x_3, x_4, x_5, x_6, x_7, 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; @@ -37556,7 +37669,7 @@ lean_inc(x_21); lean_dec(x_19); x_22 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; lean_inc(x_6); -x_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_21); +x_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_22, x_2, x_3, x_4, x_5, x_6, x_7, x_21); if (lean_obj_tag(x_23) == 0) { uint8_t x_24; @@ -37749,7 +37862,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabCond(lean_object* _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_9 = l_Lean_PrettyPrinter_Delaborator_delabCond___closed__3; x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; @@ -37933,14 +38046,14 @@ goto block_80; block_80: { lean_object* x_18; lean_object* x_19; -x_18 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__2; +x_18 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; 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_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_18, x_1, x_2, x_3, x_4, x_5, x_6, x_17); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_18, x_1, x_2, x_3, x_4, x_5, x_6, 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; @@ -37949,14 +38062,14 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; +x_22 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; 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_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_21); +x_23 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_22, x_1, x_2, x_3, x_4, x_5, x_6, x_21); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; @@ -37967,7 +38080,7 @@ lean_inc(x_25); lean_dec(x_23); x_26 = l_Lean_PrettyPrinter_Delaborator_withTypeAscription___closed__1; lean_inc(x_5); -x_27 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_26, x_1, x_2, x_3, x_4, x_5, x_6, x_25); +x_27 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_26, x_1, x_2, x_3, x_4, x_5, x_6, x_25); if (lean_obj_tag(x_27) == 0) { uint8_t x_28; @@ -38884,7 +38997,7 @@ x_23 = l_Lean_PrettyPrinter_Delaborator_delabSigmaCore___lambda__3___closed__1; x_24 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_24, 0, x_23); lean_closure_set(x_24, 1, x_22); -x_25 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +x_25 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_20); return x_25; } } @@ -38901,7 +39014,7 @@ x_11 = l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__1; x_12 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_12, 0, x_11); lean_closure_set(x_12, 1, x_10); -x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_13 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_14 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_13, x_12, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_14; } @@ -39359,7 +39472,7 @@ lean_inc(x_4); x_18 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_18, 0, x_4); lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_19 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_4); x_20 = l_Lean_Syntax_node2(x_4, x_19, x_18, x_1); lean_inc(x_4); @@ -39469,7 +39582,7 @@ lean_inc(x_4); x_19 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_19, 0, x_4); lean_ctor_set(x_19, 1, x_18); -x_20 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9; +x_20 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; lean_inc(x_4); x_21 = l_Lean_Syntax_node2(x_4, x_20, x_19, x_2); lean_inc(x_4); @@ -39479,11 +39592,11 @@ lean_inc(x_4); x_24 = lean_alloc_ctor(2, 2, 0); lean_ctor_set(x_24, 0, x_4); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6; +x_25 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8; lean_inc(x_16); lean_inc(x_4); x_26 = l_Lean_Syntax_node5(x_4, x_25, x_17, x_16, x_22, x_24, x_3); -x_27 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4; +x_27 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6; lean_inc(x_4); x_28 = l_Lean_Syntax_node1(x_4, x_27, x_26); x_29 = l_Lean_PrettyPrinter_Delaborator_delabDoElems___lambda__4___closed__2; @@ -39840,7 +39953,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__1; x_2 = l_Lean_PrettyPrinter_Delaborator_delabDoElems___closed__4; -x_3 = lean_unsigned_to_nat(1084u); +x_3 = lean_unsigned_to_nat(1157u); x_4 = lean_unsigned_to_nat(40u); x_5 = l_Lean_PrettyPrinter_Delaborator_delabFVar___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -39956,7 +40069,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_42 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_40, x_41, x_1, x_2, x_3, x_4, x_5, x_6, x_10); +x_42 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_40, x_41, x_1, x_2, x_3, x_4, x_5, x_6, x_10); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; @@ -39972,7 +40085,7 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); lean_inc(x_1); -x_46 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__6(x_45, x_41, x_1, x_2, x_3, x_4, x_5, x_6, x_44); +x_46 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__5(x_45, x_41, x_1, x_2, x_3, x_4, x_5, x_6, x_44); if (lean_obj_tag(x_46) == 0) { lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; @@ -40282,14 +40395,14 @@ goto block_121; block_121: { lean_object* x_109; lean_object* x_110; -x_109 = l_Lean_PrettyPrinter_Delaborator_delabMData___closed__1; +x_109 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1; 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_110 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__2(x_109, x_1, x_2, x_3, x_4, x_5, x_6, x_10); +x_110 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(x_109, x_1, x_2, x_3, x_4, x_5, x_6, x_10); if (lean_obj_tag(x_110) == 0) { lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; @@ -40719,7 +40832,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabDo(lean_object* x _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_9 = l_Lean_PrettyPrinter_Delaborator_delabDo___closed__2; x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; @@ -41453,7 +41566,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabNameMkStr(lean_ob _start: { lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1; +x_8 = l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1; x_9 = l_Lean_PrettyPrinter_Delaborator_delabNameMkStr___closed__2; x_10 = l_Lean_PrettyPrinter_Delaborator_whenPPOption(x_8, x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7); return x_10; @@ -41951,46 +42064,48 @@ lean_dec(x_2); return x_4; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1() { _start: { lean_object* x_1; -x_1 = l_Lean_pp_piBinderTypes; +x_1 = l_Lean_maxRecDepth; return x_1; } } -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) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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: { -uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) +uint8_t x_11; +x_11 = !lean_is_exclusive(x_8); +if (x_11 == 0) { -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; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); +x_16 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_8, x_9, x_10); +return x_16; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t 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_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); -x_26 = lean_ctor_get_uint8(x_5, sizeof(void*)*11 + 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_17 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_8, 1); +x_19 = lean_ctor_get(x_8, 3); +x_20 = lean_ctor_get(x_8, 5); +x_21 = lean_ctor_get(x_8, 6); +x_22 = lean_ctor_get(x_8, 7); +x_23 = lean_ctor_get(x_8, 8); +x_24 = lean_ctor_get(x_8, 9); +x_25 = lean_ctor_get(x_8, 10); +x_26 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +lean_inc(x_25); lean_inc(x_24); lean_inc(x_23); lean_inc(x_22); @@ -41999,30 +42114,221 @@ lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -lean_inc(x_16); +lean_dec(x_8); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1; +x_28 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_27); +x_29 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_29, 0, x_17); +lean_ctor_set(x_29, 1, x_18); +lean_ctor_set(x_29, 2, x_1); +lean_ctor_set(x_29, 3, x_19); +lean_ctor_set(x_29, 4, x_28); +lean_ctor_set(x_29, 5, x_20); +lean_ctor_set(x_29, 6, x_21); +lean_ctor_set(x_29, 7, x_22); +lean_ctor_set(x_29, 8, x_23); +lean_ctor_set(x_29, 9, x_24); +lean_ctor_set(x_29, 10, x_25); +lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*11 + 1, x_2); +x_30 = l_Lean_PrettyPrinter_Delaborator_delab(x_3, x_4, x_5, x_6, x_29, x_9, x_10); +return x_30; +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_pp_piBinderTypes; +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__3() { +_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_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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; +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; +} +} +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: +{ +lean_object* x_8; lean_object* x_9; uint8_t 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; uint8_t x_18; uint8_t x_19; +x_8 = lean_ctor_get(x_5, 2); +lean_inc(x_8); +x_9 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1; +x_10 = 1; +x_11 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_8, x_9, x_10); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2; +x_13 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_11, x_12); +x_14 = lean_st_ref_get(x_6, x_7); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_5); -x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__1; -x_28 = 1; -x_29 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_16, x_27, x_28); -x_30 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_30, 0, x_14); -lean_ctor_set(x_30, 1, x_15); -lean_ctor_set(x_30, 2, x_29); -lean_ctor_set(x_30, 3, x_17); -lean_ctor_set(x_30, 4, x_18); -lean_ctor_set(x_30, 5, x_19); -lean_ctor_set(x_30, 6, x_20); -lean_ctor_set(x_30, 7, x_21); -lean_ctor_set(x_30, 8, x_22); -lean_ctor_set(x_30, 9, x_23); -lean_ctor_set(x_30, 10, x_24); -lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_25); -lean_ctor_set_uint8(x_30, sizeof(void*)*11 + 1, x_26); -x_31 = l_Lean_PrettyPrinter_Delaborator_delab(x_1, x_2, x_3, x_4, x_30, x_6, x_7); +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_dec(x_15); +x_18 = l_Lean_Kernel_isDiagnosticsEnabled(x_17); +lean_dec(x_17); +if (x_18 == 0) +{ +if (x_13 == 0) +{ +x_19 = x_10; +goto block_47; +} +else +{ +uint8_t x_48; +x_48 = 0; +x_19 = x_48; +goto block_47; +} +} +else +{ +if (x_13 == 0) +{ +uint8_t x_49; +x_49 = 0; +x_19 = x_49; +goto block_47; +} +else +{ +x_19 = x_10; +goto block_47; +} +} +block_47: +{ +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_st_ref_take(x_6, x_16); +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = !lean_is_exclusive(x_21); +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_24 = lean_ctor_get(x_21, 0); +x_25 = lean_ctor_get(x_21, 4); +lean_dec(x_25); +x_26 = l_Lean_Kernel_enableDiag(x_24, x_13); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6; +lean_ctor_set(x_21, 4, x_27); +lean_ctor_set(x_21, 0, x_26); +x_28 = lean_st_ref_set(x_6, x_21, x_22); +x_29 = lean_ctor_get(x_28, 1); +lean_inc(x_29); +lean_dec(x_28); +x_30 = lean_box(0); +x_31 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1(x_11, x_13, x_1, x_2, x_3, x_4, x_30, x_5, x_6, x_29); return x_31; } +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; lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_32 = lean_ctor_get(x_21, 0); +x_33 = lean_ctor_get(x_21, 1); +x_34 = lean_ctor_get(x_21, 2); +x_35 = lean_ctor_get(x_21, 3); +x_36 = lean_ctor_get(x_21, 5); +x_37 = lean_ctor_get(x_21, 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_dec(x_21); +x_38 = l_Lean_Kernel_enableDiag(x_32, x_13); +x_39 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6; +x_40 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_33); +lean_ctor_set(x_40, 2, x_34); +lean_ctor_set(x_40, 3, x_35); +lean_ctor_set(x_40, 4, x_39); +lean_ctor_set(x_40, 5, x_36); +lean_ctor_set(x_40, 6, x_37); +x_41 = lean_st_ref_set(x_6, x_40, x_22); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +lean_dec(x_41); +x_43 = lean_box(0); +x_44 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1(x_11, x_13, x_1, x_2, x_3, x_4, x_43, x_5, x_6, x_42); +return x_44; +} +} +else +{ +lean_object* x_45; lean_object* x_46; +x_45 = lean_box(0); +x_46 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1(x_11, x_13, x_1, x_2, x_3, x_4, x_45, x_5, x_6, x_16); +return x_46; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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_2); +lean_dec(x_2); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___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_7); +return x_12; } } LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_shouldGroupWithNext(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -42308,77 +42614,59 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu _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; -} -} -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; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7() { +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; 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_4 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6() { _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_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_PrettyPrinter_Delaborator_delabMVar___lambda__1___closed__5; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__6; x_3 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10() { +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 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; -x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; +x_1 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; +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_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9() { _start: { lean_object* x_1; @@ -42386,7 +42674,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWith return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10() { _start: { lean_object* x_1; @@ -42394,12 +42682,12 @@ x_1 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWith return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11() { _start: { 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_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; +x_2 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10; x_3 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_3, 0, x_1); lean_closure_set(x_3, 1, x_2); @@ -42498,7 +42786,7 @@ goto block_73; else { lean_object* x_91; lean_object* x_92; -x_91 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__13; +x_91 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -42580,12 +42868,12 @@ 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_24 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; 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_28 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8; 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); @@ -42600,7 +42888,7 @@ 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; +x_36 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; 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; @@ -42628,12 +42916,12 @@ 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_50 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__7; 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_54 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__8; 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); @@ -42648,7 +42936,7 @@ 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; +x_62 = l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11; 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; @@ -43464,8 +43752,8 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu _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); +x_1 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___spec__1), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -43517,7 +43805,7 @@ lean_inc(x_17); 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); +x_19 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1(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_20; @@ -44139,9 +44427,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatur _start: { lean_object* x_12; lean_object* x_13; -x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11; +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; 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); +x_13 = l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppArg___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___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; @@ -44452,7 +44740,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu 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(1187u); +x_3 = lean_unsigned_to_nat(1260u); 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); @@ -44472,7 +44760,7 @@ static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignatu { lean_object* x_1; lean_object* x_2; 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); +x_2 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_SubExpr_withAppFn___at_Lean_PrettyPrinter_Delaborator_delabLetFun___spec__1), 8, 1); lean_closure_set(x_2, 0, x_1); return x_2; } @@ -44524,12 +44812,12 @@ lean_dec(x_16); if (lean_obj_tag(x_26) == 0) { lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_27 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_29 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_30 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_30, 0, x_29); lean_closure_set(x_30, 1, x_28); @@ -44600,12 +44888,12 @@ case 0: { 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_42 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_44 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_45 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_45, 0, x_44); lean_closure_set(x_45, 1, x_43); @@ -44669,12 +44957,12 @@ case 1: { 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_56 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_58 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_59 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_59, 0, x_58); lean_closure_set(x_59, 1, x_57); @@ -44738,12 +45026,12 @@ case 2: { 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_70 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_72 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_73 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_73, 0, x_72); lean_closure_set(x_73, 1, x_71); @@ -44807,12 +45095,12 @@ case 3: { 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_84 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_86 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_87 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_87, 0, x_86); lean_closure_set(x_87, 1, x_85); @@ -44878,7 +45166,7 @@ lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; le 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_99 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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); @@ -44948,12 +45236,12 @@ case 5: { 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_113 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_115 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_116 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_116, 0, x_115); lean_closure_set(x_116, 1, x_114); @@ -45017,12 +45305,12 @@ case 6: { 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_127 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_129 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_130 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_130, 0, x_129); lean_closure_set(x_130, 1, x_128); @@ -45086,12 +45374,12 @@ case 7: { 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_141 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_143 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_144 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_144, 0, x_143); lean_closure_set(x_144, 1, x_142); @@ -45155,12 +45443,12 @@ case 8: { 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_155 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_157 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_158 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_158, 0, x_157); lean_closure_set(x_158, 1, x_156); @@ -45224,12 +45512,12 @@ case 9: { 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_169 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_171 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_172 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_172, 0, x_171); lean_closure_set(x_172, 1, x_170); @@ -45293,12 +45581,12 @@ case 10: { 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_183 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_185 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_186 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_186, 0, x_185); lean_closure_set(x_186, 1, x_184); @@ -45362,12 +45650,12 @@ return x_196; { lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; lean_dec(x_41); -x_197 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__5; +x_197 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_199 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_200 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_200, 0, x_199); lean_closure_set(x_200, 1, x_198); @@ -45434,7 +45722,7 @@ else { 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_211 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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); @@ -45505,12 +45793,12 @@ case 1: 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_225 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_227 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_228 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_228, 0, x_227); lean_closure_set(x_228, 1, x_226); @@ -45575,12 +45863,12 @@ case 2: 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_239 = l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4; 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_241 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_242 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_242, 0, x_241); lean_closure_set(x_242, 1, x_240); @@ -45647,7 +45935,7 @@ 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_254 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__9; x_255 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_PrettyPrinter_Delaborator_delabMVar___spec__1___rarg), 9, 2); lean_closure_set(x_255, 0, x_254); lean_closure_set(x_255, 1, x_253); @@ -46272,7 +46560,71 @@ lean_dec(x_3); return x_10; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1() { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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: +{ +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; +x_12 = lean_ctor_get(x_8, 4); +lean_dec(x_12); +x_13 = lean_ctor_get(x_8, 2); +lean_dec(x_13); +x_14 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1; +x_15 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_14); +lean_ctor_set(x_8, 4, x_15); +lean_ctor_set(x_8, 2, x_1); +lean_ctor_set_uint8(x_8, sizeof(void*)*11 + 1, x_2); +x_16 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_3, x_4, x_5, x_6, x_8, x_9, x_10); +return x_16; +} +else +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; 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_29; lean_object* x_30; +x_17 = lean_ctor_get(x_8, 0); +x_18 = lean_ctor_get(x_8, 1); +x_19 = lean_ctor_get(x_8, 3); +x_20 = lean_ctor_get(x_8, 5); +x_21 = lean_ctor_get(x_8, 6); +x_22 = lean_ctor_get(x_8, 7); +x_23 = lean_ctor_get(x_8, 8); +x_24 = lean_ctor_get(x_8, 9); +x_25 = lean_ctor_get(x_8, 10); +x_26 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +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_dec(x_8); +x_27 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1; +x_28 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_27); +x_29 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_29, 0, x_17); +lean_ctor_set(x_29, 1, x_18); +lean_ctor_set(x_29, 2, x_1); +lean_ctor_set(x_29, 3, x_19); +lean_ctor_set(x_29, 4, x_28); +lean_ctor_set(x_29, 5, x_20); +lean_ctor_set(x_29, 6, x_21); +lean_ctor_set(x_29, 7, x_22); +lean_ctor_set(x_29, 8, x_23); +lean_ctor_set(x_29, 9, x_24); +lean_ctor_set(x_29, 10, x_25); +lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_26); +lean_ctor_set_uint8(x_29, sizeof(void*)*11 + 1, x_2); +x_30 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_3, x_4, x_5, x_6, x_29, x_9, x_10); +return x_30; +} +} +} +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -46280,7 +46632,7 @@ x_1 = l_Lean_pp_universes; return x_1; } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2() { _start: { lean_object* x_1; @@ -46288,73 +46640,132 @@ x_1 = l_Lean_pp_fullNames; return x_1; } } -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) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_6); -if (x_9 == 0) +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; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; uint8_t x_22; +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__1; +x_11 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_9, x_10, x_1); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2; +x_13 = 1; +x_14 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_11, x_12, x_13); +x_15 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_14, x_15); +x_17 = lean_st_ref_get(x_7, x_8); +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 0); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Kernel_isDiagnosticsEnabled(x_20); +lean_dec(x_20); +if (x_21 == 0) { -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 = 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; +if (x_16 == 0) +{ +x_22 = x_13; +goto block_50; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t 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; -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); -x_29 = lean_ctor_get_uint8(x_6, sizeof(void*)*11 + 1); -lean_inc(x_27); -lean_inc(x_26); -lean_inc(x_25); +uint8_t x_51; +x_51 = 0; +x_22 = x_51; +goto block_50; +} +} +else +{ +if (x_16 == 0) +{ +uint8_t x_52; +x_52 = 0; +x_22 = x_52; +goto block_50; +} +else +{ +x_22 = x_13; +goto block_50; +} +} +block_50: +{ +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_st_ref_take(x_7, x_19); +x_24 = lean_ctor_get(x_23, 0); 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_dec(x_6); -x_30 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__1; -x_31 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_19, x_30, x_1); -x_32 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1___closed__2; -x_33 = 1; -x_34 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_31, x_32, x_33); -x_35 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_35, 0, x_17); -lean_ctor_set(x_35, 1, x_18); -lean_ctor_set(x_35, 2, x_34); -lean_ctor_set(x_35, 3, x_20); -lean_ctor_set(x_35, 4, x_21); -lean_ctor_set(x_35, 5, x_22); -lean_ctor_set(x_35, 6, x_23); -lean_ctor_set(x_35, 7, x_24); -lean_ctor_set(x_35, 8, x_25); -lean_ctor_set(x_35, 9, x_26); -lean_ctor_set(x_35, 10, x_27); -lean_ctor_set_uint8(x_35, sizeof(void*)*11, x_28); -lean_ctor_set_uint8(x_35, sizeof(void*)*11 + 1, x_29); -x_36 = l_Lean_PrettyPrinter_Delaborator_delabConst(x_2, x_3, x_4, x_5, x_35, x_7, x_8); -return x_36; +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_27 = lean_ctor_get(x_24, 0); +x_28 = lean_ctor_get(x_24, 4); +lean_dec(x_28); +x_29 = l_Lean_Kernel_enableDiag(x_27, x_16); +x_30 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6; +lean_ctor_set(x_24, 4, x_30); +lean_ctor_set(x_24, 0, x_29); +x_31 = lean_st_ref_set(x_7, x_24, x_25); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_box(0); +x_34 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(x_14, x_16, x_2, x_3, x_4, x_5, x_33, x_6, x_7, 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; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_35 = lean_ctor_get(x_24, 0); +x_36 = lean_ctor_get(x_24, 1); +x_37 = lean_ctor_get(x_24, 2); +x_38 = lean_ctor_get(x_24, 3); +x_39 = lean_ctor_get(x_24, 5); +x_40 = lean_ctor_get(x_24, 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_dec(x_24); +x_41 = l_Lean_Kernel_enableDiag(x_35, x_16); +x_42 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6; +x_43 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_36); +lean_ctor_set(x_43, 2, x_37); +lean_ctor_set(x_43, 3, x_38); +lean_ctor_set(x_43, 4, x_42); +lean_ctor_set(x_43, 5, x_39); +lean_ctor_set(x_43, 6, x_40); +x_44 = lean_st_ref_set(x_7, x_43, x_25); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +lean_dec(x_44); +x_46 = lean_box(0); +x_47 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(x_14, x_16, x_2, x_3, x_4, x_5, x_46, x_6, x_7, x_45); +return x_47; +} +} +else +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_box(0); +x_49 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(x_14, x_16, x_2, x_3, x_4, x_5, x_48, x_6, x_7, x_19); +return x_49; +} } } } @@ -46369,7 +46780,7 @@ 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); +x_13 = lean_alloc_closure((void*)(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___boxed), 8, 1); lean_closure_set(x_13, 0, x_12); x_14 = lean_unsigned_to_nat(0u); lean_inc(x_7); @@ -46472,13 +46883,24 @@ 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) { +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, lean_object* x_9, lean_object* x_10) { +_start: +{ +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_2); +lean_dec(x_2); +x_12 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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_7); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___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_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__1(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2(x_9, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } } @@ -46808,6 +47230,14 @@ l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2___clo lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__2___closed__6); l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__3___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_appFieldNotationCandidate_x3f___lambda__3___closed__1); +l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1(); +lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__1); +l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2(); +lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__2); +l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3(); +lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__3); +l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4(); +lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_stripParentProjections___closed__4); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__2(); @@ -46822,10 +47252,6 @@ l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__6 = lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__4___closed__6); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__6___closed__3); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___lambda__8___closed__2(); @@ -46842,50 +47268,42 @@ l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__4 = _init_l_Lean lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__4); l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__5(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppExplicitCore___closed__5); +l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg = _init_l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_instInhabitedAppImplicitArg); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkNamedArg___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__3); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__4); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__5); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__6); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArgStx___closed__7); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_mkArg___closed__4); l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1 = _init_l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1(); lean_mark_persistent(l_List_firstM___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryUnexpand___spec__1___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders_go___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___lambda__2___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__2); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore_tryAppUnexpanders___closed__3); -l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1 = _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1(); -lean_mark_persistent(l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___spec__6___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__2___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__5___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__3); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__6___closed__4); -l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__7___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__9___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___lambda__10___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppImplicitCore___closed__2(); @@ -46917,25 +47335,19 @@ lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp___cl if (builtin) {res = l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabApp(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_PrettyPrinter_Delaborator_AppMatchState_params___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_params___default(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_params___default); -l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default(); +}l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_motive___default); l_Lean_PrettyPrinter_Delaborator_AppMatchState_motiveNamed___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_motiveNamed___default(); l_Lean_PrettyPrinter_Delaborator_AppMatchState_discrs___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_discrs___default(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_discrs___default); +l_Lean_PrettyPrinter_Delaborator_AppMatchState_hNames_x3f___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_hNames_x3f___default(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_hNames_x3f___default); +l_Lean_PrettyPrinter_Delaborator_AppMatchState_alts___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_alts___default(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_alts___default); l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_varNames___default); l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_rhss___default); -l_Lean_PrettyPrinter_Delaborator_AppMatchState_moreArgs___default = _init_l_Lean_PrettyPrinter_Delaborator_AppMatchState_moreArgs___default(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_AppMatchState_moreArgs___default); -l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns_usingNamesAux___spec__1___rarg___closed__1); -l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__1 = _init_l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__1(); -lean_mark_persistent(l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__1); -l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2 = _init_l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2(); -lean_mark_persistent(l_Array_mapIdxM_map___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_delabPatterns___spec__5___closed__2); l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1(); lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__1); l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__2(); @@ -46946,6 +47358,8 @@ l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delabo lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__4); l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5 = _init_l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5(); lean_mark_persistent(l___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop___rarg___lambda__1___closed__5); +l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at___private_Lean_PrettyPrinter_Delaborator_Builtins_0__Lean_PrettyPrinter_Delaborator_skippingBinders_loop_visitLambda___spec__1___rarg___closed__1); l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1 = _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1(); lean_mark_persistent(l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__1); l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2 = _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__2(); @@ -46954,24 +47368,32 @@ l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2_ lean_mark_persistent(l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__3); l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4 = _init_l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4(); lean_mark_persistent(l_Lean_getConstInfo___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__2___closed__4); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__1); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__2); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__3); -l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__10___closed__4); -l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1 = _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1(); -lean_mark_persistent(l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__12___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__3); -l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__4___closed__4); +l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1 = _init_l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__1); +l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2 = _init_l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__18___closed__2); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__1); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__19___closed__2); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__1); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__2); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__3); +l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4 = _init_l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__21___closed__4); +l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1 = _init_l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1(); +lean_mark_persistent(l_panic___at_Lean_PrettyPrinter_Delaborator_delabAppMatch___spec__23___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__3___closed__4); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__2(); @@ -46994,6 +47416,12 @@ l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10 = _init_ lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__10); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__11); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__6___closed__12); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___lambda__7___closed__2); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2(); @@ -47002,8 +47430,6 @@ l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3 = _init_l_Lean_Pretty lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__3); l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__4); -l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__5); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__1); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch___closed__2(); @@ -47013,28 +47439,32 @@ lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch if (builtin) {res = l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabAppMatch(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__1); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__2); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__3); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__4); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__5); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__6); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__7); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__8 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__8(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__8); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__9); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__10); -l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__1___closed__11); +}l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__4); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__5); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__6); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__7); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__8); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__9); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__10); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__11); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__12); +l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___lambda__2___closed__13); l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__1); l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabLetFun___closed__2(); @@ -47062,8 +47492,6 @@ l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3 = _init_l_Lean_PrettyPri lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabMData___closed__3); l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabMData___closed__4); -l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabMData___closed__5); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__1 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__1(); lean_mark_persistent(l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__1); l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__2 = _init_l___regBuiltin_Lean_PrettyPrinter_Delaborator_delabMData___closed__2(); @@ -47710,8 +48138,20 @@ 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_delabTy___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___lambda__1___closed__1); 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_delabTy___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__2); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__3 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__3(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__3); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__4 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__4(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__4); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__5); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabTy___closed__6); 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(); @@ -47734,10 +48174,6 @@ l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__1 lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__10); l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__11); -l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParams___closed__12(); -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_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(); @@ -47758,10 +48194,10 @@ l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature_delabParamsAux___closed 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___lambda__2___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__1); +l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2 = _init_l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_delabConstWithSignature___lambda__2___closed__2); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c index cc7e1e7cf097..43c8d89a6f2c 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/FieldNotation.c @@ -25,6 +25,7 @@ lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Name_isAnonymous(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -54,14 +55,13 @@ uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__2___closed__1; lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_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_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___lambda__2___closed__2; static lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___closed__2; extern lean_object* l_Lean_instInhabitedExpr; lean_object* l_Lean_FVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_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* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -69,35 +69,37 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f_ lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___closed__2; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__4(uint8_t, 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_instInhabitedProjectionFunctionInfo; LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_typeMatchesBaseName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_BinderInfo_isExplicit(uint8_t); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FVarId_getUserName(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_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); +static lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__2___closed__2; extern lean_object* l_Lean_ppNoDotAttr; -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(uint8_t, lean_object*, lean_object*, 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_PrettyPrinter_Delaborator_fieldNotationCandidate_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_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__2___boxed(lean_object**); 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_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); @@ -2436,7 +2438,7 @@ x_10 = l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f(x_1, x_2, x_9 return x_10; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; uint8_t x_9; @@ -2474,20 +2476,19 @@ return x_15; } } } -static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1() { +static lean_object* _init_l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1() { _start: { -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; +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_box(x_1); -x_4 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; +x_3 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___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; lean_object* x_10; lean_object* x_14; lean_object* x_15; lean_object* x_16; @@ -2501,124 +2502,129 @@ lean_dec(x_14); x_24 = l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo(x_23, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_24) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +lean_object* x_25; 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_42; x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); -lean_dec(x_25); -x_27 = lean_ctor_get(x_26, 1); +x_27 = lean_ctor_get(x_24, 1); 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_dec(x_26); -x_29 = lean_ctor_get(x_24, 1); -lean_inc(x_29); lean_dec(x_24); -x_30 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1; -x_9 = x_30; -x_10 = x_29; +x_28 = lean_ctor_get(x_25, 0); +lean_inc(x_28); +lean_dec(x_25); +x_29 = lean_ctor_get(x_26, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_26, 1); +lean_inc(x_30); +lean_dec(x_26); +x_42 = lean_unbox(x_30); +lean_dec(x_30); +if (x_42 == 0) +{ +lean_object* x_43; +lean_dec(x_29); +lean_dec(x_28); +x_43 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; +x_9 = x_43; +x_10 = x_27; goto block_13; } else { if (x_2 == 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; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_31 = lean_ctor_get(x_24, 1); -lean_inc(x_31); -lean_dec(x_24); -x_32 = lean_ctor_get(x_26, 0); -lean_inc(x_32); -lean_dec(x_26); -x_33 = lean_unsigned_to_nat(0u); -x_34 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_33); -x_35 = lean_unsigned_to_nat(1u); -x_36 = lean_nat_add(x_32, x_35); -lean_dec(x_32); -x_37 = lean_nat_dec_eq(x_34, x_36); -lean_dec(x_36); -lean_dec(x_34); -x_38 = lean_box(0); -x_39 = lean_box(x_37); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_9 = x_40; -x_10 = x_31; -goto block_13; +lean_object* x_44; +x_44 = lean_box(0); +x_31 = x_44; +goto block_41; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; -x_41 = lean_ctor_get(x_24, 1); -lean_inc(x_41); -lean_dec(x_24); -x_42 = lean_ctor_get(x_26, 0); -lean_inc(x_42); -lean_dec(x_26); -x_43 = lean_unsigned_to_nat(0u); -x_44 = lean_nat_dec_eq(x_42, x_43); -if (x_44 == 0) +lean_object* x_45; uint8_t x_46; +x_45 = lean_unsigned_to_nat(0u); +x_46 = lean_nat_dec_eq(x_29, x_45); +if (x_46 == 0) { -lean_object* x_45; -lean_dec(x_42); -x_45 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1; -x_9 = x_45; -x_10 = x_41; +lean_object* x_47; +lean_dec(x_29); +lean_dec(x_28); +x_47 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; +x_9 = x_47; +x_10 = x_27; goto block_13; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; uint8_t x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_46 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_43); -x_47 = lean_unsigned_to_nat(1u); -x_48 = lean_nat_add(x_42, x_47); -lean_dec(x_42); -x_49 = lean_nat_dec_eq(x_46, x_48); -lean_dec(x_48); -lean_dec(x_46); -x_50 = lean_box(0); -x_51 = lean_box(x_49); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_51); -lean_ctor_set(x_52, 1, x_50); -x_9 = x_52; -x_10 = x_41; +lean_object* x_48; +x_48 = lean_box(0); +x_31 = x_48; +goto block_41; +} +} +} +block_41: +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; +lean_dec(x_31); +x_32 = lean_unsigned_to_nat(0u); +x_33 = l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(x_1, x_32); +x_34 = lean_unsigned_to_nat(1u); +x_35 = lean_nat_add(x_29, x_34); +lean_dec(x_29); +x_36 = lean_nat_dec_eq(x_33, x_35); +lean_dec(x_35); +lean_dec(x_33); +if (x_36 == 0) +{ +lean_object* x_37; +lean_dec(x_28); +x_37 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; +x_9 = x_37; +x_10 = x_27; goto block_13; } +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_38, 0, x_28); +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_9 = x_40; +x_10 = x_27; +goto block_13; } } } else { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_24, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_24, 1); -lean_inc(x_54); +lean_object* x_49; lean_object* x_50; +x_49 = lean_ctor_get(x_24, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_24, 1); +lean_inc(x_50); lean_dec(x_24); -x_15 = x_53; -x_16 = x_54; +x_15 = x_49; +x_16 = x_50; goto block_22; } } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_dec(x_14); -x_55 = l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___closed__2; -x_56 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(x_55, x_4, x_5, x_6, x_7, x_8); -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_15 = x_57; -x_16 = x_58; +x_51 = l___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_projInfo___closed__2; +x_52 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1(x_51, x_4, x_5, x_6, x_7, x_8); +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_15 = x_53; +x_16 = x_54; goto block_22; } block_13: @@ -2640,7 +2646,7 @@ if (x_17 == 0) { lean_object* x_18; lean_dec(x_15); -x_18 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1; +x_18 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; x_9 = x_18; x_10 = x_16; goto block_13; @@ -2661,7 +2667,7 @@ else { lean_object* x_21; lean_dec(x_15); -x_21 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1; +x_21 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1; x_9 = x_21; x_10 = x_16; goto block_13; @@ -2670,45 +2676,44 @@ goto block_13; } } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj(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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; x_8 = l_Lean_Expr_isApp(x_2); if (x_8 == 0) { -uint8_t x_9; lean_object* x_10; lean_object* x_11; +lean_object* x_9; lean_object* x_10; lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_9 = 0; -x_10 = lean_box(x_9); -x_11 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_7); -return x_11; +x_9 = lean_box(0); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_7); +return x_10; } else { -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1(x_2, x_1, x_12, x_3, x_4, x_5, x_6, x_7); +lean_object* x_11; lean_object* x_12; +x_11 = lean_box(0); +x_12 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1(x_2, x_1, x_11, 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_13; +return x_12; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___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_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; -x_7 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_isParentProj___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_throwError___at_Lean_PrettyPrinter_Delaborator_parentProj_x3f___spec__1(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); @@ -2716,13 +2721,13 @@ lean_dec(x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { uint8_t x_9; lean_object* x_10; x_9 = lean_unbox(x_2); lean_dec(x_2); -x_10 = l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1(x_1, x_9, x_3, x_4, x_5, x_6, x_7, x_8); +x_10 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___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); @@ -2732,6 +2737,108 @@ lean_dec(x_1); return x_10; } } +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_parentProj_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_1); +lean_dec(x_1); +x_9 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(x_8, x_2, x_3, x_4, x_5, x_6, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l_Lean_PrettyPrinter_Delaborator_parentProj_x3f(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +if (lean_obj_tag(x_8) == 0) +{ +lean_object* x_9; +x_9 = lean_ctor_get(x_8, 0); +lean_inc(x_9); +if (lean_obj_tag(x_9) == 0) +{ +uint8_t x_10; +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) +{ +lean_object* x_11; uint8_t x_12; lean_object* x_13; +x_11 = lean_ctor_get(x_8, 0); +lean_dec(x_11); +x_12 = 0; +x_13 = lean_box(x_12); +lean_ctor_set(x_8, 0, x_13); +return x_8; +} +else +{ +lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_8, 1); +lean_inc(x_14); +lean_dec(x_8); +x_15 = 0; +x_16 = lean_box(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_14); +return x_17; +} +} +else +{ +uint8_t x_18; +lean_dec(x_9); +x_18 = !lean_is_exclusive(x_8); +if (x_18 == 0) +{ +lean_object* x_19; uint8_t x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_8, 0); +lean_dec(x_19); +x_20 = 1; +x_21 = lean_box(x_20); +lean_ctor_set(x_8, 0, x_21); +return x_8; +} +else +{ +lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; +x_22 = lean_ctor_get(x_8, 1); +lean_inc(x_22); +lean_dec(x_8); +x_23 = 1; +x_24 = lean_box(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; +} +} +} +else +{ +uint8_t x_26; +x_26 = !lean_is_exclusive(x_8); +if (x_26 == 0) +{ +return x_8; +} +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_8, 0); +x_28 = lean_ctor_get(x_8, 1); +lean_inc(x_28); +lean_inc(x_27); +lean_dec(x_8); +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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isParentProj___boxed(lean_object* x_1, 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: { @@ -2787,8 +2894,8 @@ l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__3___closed lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__3___closed__1); l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__4___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__4___closed__1(); lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f___lambda__4___closed__1); -l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_isParentProj___lambda__1___closed__1); +l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1 = _init_l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_PrettyPrinter_Delaborator_parentProj_x3f___lambda__1___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c index 987be743e545..10ac19b0d371 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Delaborator/TopDownAnalyze.c @@ -19,10 +19,8 @@ static lean_object* l_Lean_getPPAnalysisLetVarType___closed__2; static lean_object* l_Lean_getPPAnalysisHole___closed__2; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_inspectOutParams___closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__1___closed__4; -static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__5___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectTrivialBottomUps___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_pp_analyze_knowsType; static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__18; @@ -45,6 +43,7 @@ extern lean_object* l_Lean_profiler; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__19; LEAN_EXPORT lean_object* l_Lean_pp_analyze_omitMax; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_86____closed__1; +extern lean_object* l_Lean_trace_profiler_useHearbeats; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__10; lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -63,6 +62,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__13; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeProj___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectHigherOrders___spec__1___boxed(lean_object**); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__7; static lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__2; @@ -98,6 +98,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instInh static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze___closed__1; uint8_t l_Lean_Level_hasParam(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2___closed__1; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp(lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isRawNatLit(lean_object*); @@ -109,9 +110,9 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replace LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody_x27___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLam___spec__3(lean_object*); uint8_t l_Lean_Expr_isApp(lean_object*); +static double l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__5; static lean_object* l_Lean_getPPAnalysisNamedArg___closed__1; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_47____closed__4; -double l_Lean_trace_profiler_threshold_getSecs(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isSimpleHOFun(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_8____closed__8; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_State_postponed___default; @@ -145,13 +146,13 @@ LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnal static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_320____closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_inspectOutParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withLetBody___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_getPPFieldNotation(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__20(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLam___closed__1; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__4; lean_object* l_instInhabitedReaderT___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -178,13 +179,12 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isNonConstFun(lean_obj static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__44; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___spec__3___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalyzeTypeAscriptions___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_getPPAnalysisSkip___closed__1; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_47____closed__1; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instMonadReaderOfSubExprAnalyzeM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeProj(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -193,7 +193,7 @@ lean_object* l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getPPAnalysisNeedsType___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_3467____closed__2; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeApp___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__4; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__38(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -201,6 +201,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasLeve LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkpointDefEq___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic___spec__1(lean_object*); lean_object* l_Lean_Expr_ReplaceLevelImpl_replaceUnsafe(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_getPPAnalysisNamedArg(lean_object*); @@ -236,7 +237,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic_core___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_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__5___boxed(lean_object**); uint8_t l_Lean_Expr_isAtomic(lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWithVars___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PrettyPrinter_Delaborator_fieldNotationCandidate_x3f(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -246,6 +247,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delabora extern lean_object* l_instInhabitedPUnit; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__5___lambda__5___boxed(lean_object**); static lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkpointDefEq___spec__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__20; uint8_t l_Lean_Expr_hasMVar(lean_object*); @@ -288,12 +290,10 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_mvarNam LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getPPAnalysisNamedArg___closed__2; LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeLet___spec__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_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__5(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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withBindingBody___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_applyFunBinderHeuristic_core___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(lean_object*, lean_object*); lean_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_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___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*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_203____closed__5; LEAN_EXPORT uint8_t l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasLevelMVarAtCurrDepth___lambda__1(lean_object*, lean_object*); @@ -315,8 +315,10 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPr LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeSort___rarg(lean_object*); static lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeConst___spec__1___closed__2; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_398____closed__5; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__21; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4(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*); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(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*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerInternalExceptionId(lean_object*, lean_object*); @@ -346,7 +348,7 @@ LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_PrettyPrinter_Delaborator_Top LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_125____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern uint8_t l_Lean_instInhabitedBinderInfo; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___lambda__1___boxed(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__23; @@ -450,6 +452,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPr static lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__4; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_getPPAnalysisHole___boxed(lean_object*); +extern lean_object* l_Lean_trace_profiler_threshold; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_47____closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_tryUnify___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_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_359_(lean_object*); @@ -475,6 +478,7 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_replaceLPsWi LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withType___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_maybeSetExplicit___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_3467____closed__1; static lean_object* l_Lean_PrettyPrinter_Delaborator_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_10698____closed__2; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_164____closed__4; extern lean_object* l_Lean_Meta_instMonadMetaM; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_320_(lean_object*); @@ -547,7 +551,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_L static lean_object* l_Lean_PrettyPrinter_Delaborator_topDownAnalyze___lambda__3___closed__1; static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_8____closed__3; lean_object* l_StateT_lift___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyze_analyzeMData___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getExpr___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_hBinOpHeuristic___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__1; @@ -621,6 +624,7 @@ LEAN_EXPORT uint8_t l_Lean_getPPAnalysisHole(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_437____closed__4; lean_object* l_Array_shrink___rarg(lean_object*, lean_object*); lean_object* l_Lean_FindLevelMVar_main(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_withNaryArg___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_398____closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_collectBottomUps___spec__1___boxed(lean_object**); @@ -812,6 +816,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_hasMVar static lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_320____closed__2; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_tryUnify(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___spec__3___rarg___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_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__3; static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_annotateBoolAt___closed__2; static lean_object* l_Lean_Meta_checkpointDefEq___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkpointDefEq___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_getPos___at_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg___spec__3(lean_object*); @@ -825,7 +830,7 @@ lean_object* l_Lean_Meta_isConstructorApp_x3f(lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_analyzeAppStagedCore_analyzeArg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_setPostponed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_isFOLike(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_returnsPi___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -871,7 +876,6 @@ static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___l LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_PrettyPrinter_Delaborator_TopDownAnalyze___hyg_86_(lean_object*); static lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_isHBinOp___lambda__1___closed__11; lean_object* l_Lean_Meta_processPostponed(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_instMonadReaderOfSubExprAnalyzeM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_TopDownAnalyze_checkpointDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -4830,7 +4834,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(212u); +x_3 = lean_unsigned_to_nat(213u); 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); @@ -31407,119 +31411,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at_Lean_PrettyPr return x_2; } } -static double _init_l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_io_mono_nanos_now(x_6); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_io_mono_nanos_now(x_12); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; uint8_t x_16; lean_object* x_17; double x_18; double x_19; double x_20; double x_21; double x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_15 = lean_ctor_get(x_13, 0); -x_16 = 0; -x_17 = lean_unsigned_to_nat(0u); -x_18 = l_Float_ofScientific(x_8, x_16, x_17); -lean_dec(x_8); -x_19 = l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; -x_20 = lean_float_div(x_18, x_19); -x_21 = l_Float_ofScientific(x_15, x_16, x_17); -lean_dec(x_15); -x_22 = lean_float_div(x_21, x_19); -x_23 = lean_box_float(x_20); -x_24 = lean_box_float(x_22); -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, 2, 0); -lean_ctor_set(x_26, 0, x_11); -lean_ctor_set(x_26, 1, x_25); -lean_ctor_set(x_13, 0, x_26); -return x_13; -} -else -{ -lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; double x_31; double x_32; double x_33; double x_34; double x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_dec(x_13); -x_29 = 0; -x_30 = lean_unsigned_to_nat(0u); -x_31 = l_Float_ofScientific(x_8, x_29, x_30); -lean_dec(x_8); -x_32 = l_Lean_withStartStopSeconds___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3___closed__1; -x_33 = lean_float_div(x_31, x_32); -x_34 = l_Float_ofScientific(x_27, x_29, x_30); -lean_dec(x_27); -x_35 = lean_float_div(x_34, x_32); -x_36 = lean_box_float(x_33); -x_37 = lean_box_float(x_35); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_11); -lean_ctor_set(x_39, 1, 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_28); -return x_40; -} -} -else -{ -uint8_t x_41; -lean_dec(x_8); -x_41 = !lean_is_exclusive(x_10); -if (x_41 == 0) -{ -return x_10; -} -else -{ -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_10, 0); -x_43 = lean_ctor_get(x_10, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_10); -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_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { if (lean_obj_tag(x_1) == 0) @@ -31544,89 +31436,23 @@ return x_10; } } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = lean_apply_5(x_1, x_2, x_3, x_4, x_5, x_6); -if (lean_obj_tag(x_7) == 0) -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; -x_9 = lean_ctor_get(x_7, 0); -x_10 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_10, 0, x_9); -lean_ctor_set(x_7, 0, x_10); -return x_7; -} -else -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_7, 0); -x_12 = lean_ctor_get(x_7, 1); -lean_inc(x_12); -lean_inc(x_11); -lean_dec(x_7); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); -return x_14; -} -} -else -{ -uint8_t x_15; -x_15 = !lean_is_exclusive(x_7); -if (x_15 == 0) -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_7, 0); -x_17 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set_tag(x_7, 0); -lean_ctor_set(x_7, 0, x_17); -return x_7; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_18 = lean_ctor_get(x_7, 0); -x_19 = lean_ctor_get(x_7, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_7); -x_20 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_20, 0, x_18); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_19); -return x_21; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_dec(x_6); lean_inc(x_9); -x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__4(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Meta_processPostponed___spec__3(x_1, x_5, x_2, x_3, x_7, x_8, x_9, x_10, x_11); x_13 = lean_ctor_get(x_12, 1); lean_inc(x_13); lean_dec(x_12); -x_14 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__4(x_4, x_7, x_8, x_9, x_10, x_13); +x_14 = l_MonadExcept_ofExcept___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__3(x_4, x_7, x_8, x_9, x_10, x_13); lean_dec(x_9); lean_dec(x_4); return x_14; } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2___closed__1() { _start: { lean_object* x_1; @@ -31634,7 +31460,7 @@ x_1 = l_Lean_profiler; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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, uint8_t x_8, double x_9, double 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_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___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, uint8_t x_8, double x_9, double 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: { double x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; @@ -31647,7 +31473,7 @@ lean_ctor_set(x_18, 1, x_3); lean_ctor_set_float(x_18, sizeof(void*)*2, x_17); lean_ctor_set_float(x_18, sizeof(void*)*2 + 8, x_17); lean_ctor_set_uint8(x_18, sizeof(void*)*2 + 16, x_2); -x_19 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1; +x_19 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2___closed__1; x_20 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_19); if (x_20 == 0) { @@ -31657,7 +31483,7 @@ lean_object* x_21; lean_object* x_22; lean_dec(x_3); lean_dec(x_1); x_21 = lean_box(0); -x_22 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); +x_22 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_18, x_21, x_12, x_13, x_14, x_15, x_16); return x_22; } else @@ -31671,7 +31497,7 @@ lean_ctor_set_float(x_23, sizeof(void*)*2, x_9); lean_ctor_set_float(x_23, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_23, sizeof(void*)*2 + 16, x_2); x_24 = lean_box(0); -x_25 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); +x_25 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_23, x_24, x_12, x_13, x_14, x_15, x_16); return x_25; } } @@ -31686,12 +31512,12 @@ lean_ctor_set_float(x_26, sizeof(void*)*2, x_9); lean_ctor_set_float(x_26, sizeof(void*)*2 + 8, x_10); lean_ctor_set_uint8(x_26, sizeof(void*)*2 + 16, x_2); x_27 = lean_box(0); -x_28 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__2(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); +x_28 = l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__1(x_4, x_5, x_11, x_6, x_26, x_27, x_12, x_13, x_14, x_15, x_16); return x_28; } } } -static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__4___closed__1() { +static lean_object* _init_l_Lean_withTraceNode___at_Lean_PrettyPrinter_Delaborator_topDownAnalyze___spec__2___lambda__3___closed__1() { _start: { lean_object* x_1; @@ -31699,16 +31525,16 @@ x_1 = lean_mk_string_from_bytes(" Date: Mon, 6 May 2024 15:55:07 -0700 Subject: [PATCH 07/44] chore: test for issue #4064 closes #4064 --- tests/lean/run/4064.lean | 11 +++++++++++ 1 file changed, 11 insertions(+) create mode 100644 tests/lean/run/4064.lean diff --git a/tests/lean/run/4064.lean b/tests/lean/run/4064.lean new file mode 100644 index 000000000000..8a2ee3b627c3 --- /dev/null +++ b/tests/lean/run/4064.lean @@ -0,0 +1,11 @@ +import Lean + +def test : Lean.CoreM (List Lean.Name) := do + let .thmInfo tval ← Lean.getConstInfo `And.left | unreachable! + return tval.all + +/-- +info: [`And.left] +-/ +#guard_msgs in +#eval test From 883a3e752d78cf0d30628817379a6001252b5595 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Tue, 7 May 2024 11:44:56 +1000 Subject: [PATCH 08/44] chore: allow omega to use classicality, in case Decidable instances are too big (#4073) From bug report at https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/omega.20regression.20in.204.2E8.2E0-rc1/near/437150155 --- src/Lean/Elab/Tactic/Omega/Frontend.lean | 1 - tests/lean/run/omega.lean | 6 +++++- 2 files changed, 5 insertions(+), 2 deletions(-) diff --git a/src/Lean/Elab/Tactic/Omega/Frontend.lean b/src/Lean/Elab/Tactic/Omega/Frontend.lean index a382690a9cdb..fe2bbd733cbb 100644 --- a/src/Lean/Elab/Tactic/Omega/Frontend.lean +++ b/src/Lean/Elab/Tactic/Omega/Frontend.lean @@ -666,7 +666,6 @@ open Lean Elab Tactic Parser.Tactic def omegaTactic (cfg : OmegaConfig) : TacticM Unit := do liftMetaFinishingTactic fun g => do let g ← g.falseOrByContra - (useClassical := false) -- because all the hypotheses we can make use of are decidable g.withContext do let hyps := (← getLocalHyps).toList trace[omega] "analyzing {hyps.length} hypotheses:\n{← hyps.mapM inferType}" diff --git a/tests/lean/run/omega.lean b/tests/lean/run/omega.lean index 73d029b6a3ef..5c6b59331fe3 100644 --- a/tests/lean/run/omega.lean +++ b/tests/lean/run/omega.lean @@ -377,8 +377,12 @@ example (i j : Nat) (p : i ≥ j) : True := by -- From https://leanprover.zulipchat.com/#narrow/stream/217875-Is-there-code-for-X.3F/topic/Nat.2Emul_sub_mod/near/428107094 example (a b : Nat) (h : a % b + 1 = 0) : False := by omega -/-! ### Fin -/ +-- From https://leanprover.zulipchat.com/#narrow/stream/270676-lean4/topic/omega.20regression.20in.204.2E8.2E0-rc1/near/437150155 +example (x : Nat) : x < 2 → + (0 = 0 → 0 = 0 → 0 = 0 → 0 = 0 → x < 2) ∧ (0 = 0 → 0 = 0 → 0 = 0 → 0 = 0 → x < 2 → x < 3) := by + omega +/-! ### Fin -/ -- Test `<` example (n : Nat) (i j : Fin n) (h : i < j) : (i : Nat) < n - 1 := by omega From bb7e6e476967a35c6b319aa110c36b7e39752b24 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Tue, 7 May 2024 17:14:09 +0200 Subject: [PATCH 09/44] chore: lake: tweak hovers for `family_def` et al (#4088) This makes the `leanArts` in `library_data leanArts : BuildJob Unit` get a hover for the generated axiom. It also simplifies the `quoteFrom` function so that it delaborates properly by using a name literal (which elaborates to `mkStr1`, `mkStr2` etc) instead of a `mkStr` application. --- src/lake/Lake/Build/Data.lean | 2 +- src/lake/Lake/Util/Family.lean | 2 +- src/lake/Lake/Util/Name.lean | 6 ++---- 3 files changed, 4 insertions(+), 6 deletions(-) diff --git a/src/lake/Lake/Build/Data.lean b/src/lake/Lake/Build/Data.lean index d376804d47a8..799f105e23bc 100644 --- a/src/lake/Lake/Build/Data.lean +++ b/src/lake/Lake/Build/Data.lean @@ -112,7 +112,7 @@ scoped macro (name := libraryDataDecl) doc?:optional(Parser.Command.docComment) "library_data " id:ident " : " ty:term : command => do let dty := mkCIdentFrom (← getRef) ``TargetData let key := Name.quoteFrom id id.getId - let id := mkIdentFrom id <| id.getId.modifyBase (`leanLib ++ ·) + let id := mkIdentFrom id (canonical := true) <| id.getId.modifyBase (`leanLib ++ ·) `($[$doc?]? family_def $id : $dty (`leanLib ++ $key) := $ty) /-- Macro for declaring new `TargetData`. -/ diff --git a/src/lake/Lake/Util/Family.lean b/src/lake/Lake/Util/Family.lean index 07926f0939db..8fab36261ac7 100644 --- a/src/lake/Lake/Util/Family.lean +++ b/src/lake/Lake/Util/Family.lean @@ -171,7 +171,7 @@ scoped macro (name := familyDef) doc?:optional(Parser.Command.docComment) let tid := extractMacroScopes fam.getId |>.name if let (tid, _) :: _ ← Macro.resolveGlobalName tid then let app := Syntax.mkApp fam #[key] - let axm := mkIdentFrom fam <| `_root_ ++ tid ++ id.getId + let axm := mkIdentFrom id (canonical := true) <| `_root_ ++ tid ++ id.getId `($[$doc?]? @[simp] axiom $axm : $app = $ty instance : FamilyDef $fam $key $ty := ⟨$axm⟩) else diff --git a/src/lake/Lake/Util/Name.lean b/src/lake/Lake/Util/Name.lean index 437be1dcc166..ad8f43387814 100644 --- a/src/lake/Lake/Util/Name.lean +++ b/src/lake/Lake/Util/Name.lean @@ -90,7 +90,5 @@ instance : LawfulCmpEq Name Name.quickCmp where open Syntax -def quoteFrom (ref : Syntax) : Name → Term -| .anonymous => mkCIdentFrom ref ``anonymous -| .str p s => mkApp (mkCIdentFrom ref ``mkStr) #[quoteFrom ref p, quote s] -| .num p v => mkApp (mkCIdentFrom ref ``mkNum) #[quoteFrom ref p, quote v] +def quoteFrom (ref : Syntax) (n : Name) : Term := + ⟨copyHeadTailInfoFrom (quote n : Term) ref⟩ From 93c06c055264dd1d5f62d89242c63b4e45e64826 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 May 2024 00:08:32 +0200 Subject: [PATCH 10/44] feat: relaxed reset/reuse in the code generator (#4100) closes #4089 --- src/Lean/Compiler/IR/ResetReuse.lean | 88 +++++++++++++++++++++------- tests/lean/4089.lean | 12 ++++ tests/lean/4089.lean.expected.out | 38 ++++++++++++ 3 files changed, 116 insertions(+), 22 deletions(-) create mode 100644 tests/lean/4089.lean create mode 100644 tests/lean/4089.lean.expected.out diff --git a/src/Lean/Compiler/IR/ResetReuse.lean b/src/Lean/Compiler/IR/ResetReuse.lean index 755f194978aa..64b78f013eee 100644 --- a/src/Lean/Compiler/IR/ResetReuse.lean +++ b/src/Lean/Compiler/IR/ResetReuse.lean @@ -29,43 +29,50 @@ Here are the main differences: does not occur in a function body. See example at `livevars.lean`. -/ -private def mayReuse (c₁ c₂ : CtorInfo) : Bool := +private def mayReuse (c₁ c₂ : CtorInfo) (relaxedReuse : Bool) : Bool := c₁.size == c₂.size && c₁.usize == c₂.usize && c₁.ssize == c₂.ssize && /- The following condition is a heuristic. - We don't want to reuse cells from different types even when they are compatible + If `relaxedReuse := false`, then we don't want to reuse cells from + different constructors even when they are compatible because it produces counterintuitive behavior. -/ - c₁.name.getPrefix == c₂.name.getPrefix + (relaxedReuse || c₁.name.getPrefix == c₂.name.getPrefix) /-- Replace `ctor` applications with `reuse` applications if compatible. `w` contains the "memory cell" being reused. -/ -private partial def S (w : VarId) (c : CtorInfo) : FnBody → FnBody +private partial def S (w : VarId) (c : CtorInfo) (relaxedReuse : Bool) (b : FnBody) : FnBody := + go b +where + go : FnBody → FnBody | .vdecl x t v@(.ctor c' ys) b => - if mayReuse c c' then + if mayReuse c c' relaxedReuse then let updtCidx := c.cidx != c'.cidx .vdecl x t (.reuse w c' updtCidx ys) b else - .vdecl x t v (S w c b) + .vdecl x t v (go b) | .jdecl j ys v b => - let v' := S w c v + let v' := go v if v == v' then - .jdecl j ys v (S w c b) + .jdecl j ys v (go b) else .jdecl j ys v' b | .case tid x xType alts => - .case tid x xType <| alts.map fun alt => alt.modifyBody (S w c) + .case tid x xType <| alts.map fun alt => alt.modifyBody go | b => if b.isTerminal then b - else let - (instr, b) := b.split - instr.setBody (S w c b) + else + let (instr, b) := b.split + instr.setBody (go b) structure Context where lctx : LocalContext := {} /-- - Contains all variables in `cases` statements in the current path. + Contains all variables in `cases` statements in the current path + and variables that are already in `reset` statements when we + invoke `R`. + We use this information to prevent double-reset in code such as ``` case x_i : obj of @@ -74,8 +81,18 @@ structure Context where Prod.mk → ... ``` + + A variable can already be in a `reset` statement when we + invoke `R` because we execute it with and without `relaxedReuse`. -/ - casesVars : PHashSet VarId := {} + alreadyFound : PHashSet VarId := {} + /-- + If `relaxedReuse := true`, then allow memory cells from different + constructors to be reused. For example, we can reuse a `PSigma.mk` + to allocate a `Prod.mk`. To avoid counterintuitive behavior, + we first try `relaxedReuse := false`, and then `relaxedReuse := true`. + -/ + relaxedReuse : Bool := false /-- We use `Context` to track join points in scope. -/ abbrev M := ReaderT Context (StateT Index Id) @@ -90,7 +107,7 @@ to replace a `ctor` withe `reuse` in `b`. -/ private def tryS (x : VarId) (c : CtorInfo) (b : FnBody) : M FnBody := do let w ← mkFresh - let b' := S w c b + let b' := S w c (← read).relaxedReuse b if b == b' then return b else @@ -102,8 +119,8 @@ private def Dfinalize (x : VarId) (c : CtorInfo) : FnBody × Bool → M FnBody private def argsContainsVar (ys : Array Arg) (x : VarId) : Bool := ys.any fun arg => match arg with - | Arg.var y => x == y - | _ => false + | .var y => x == y + | _ => false private def isCtorUsing (b : FnBody) (x : VarId) : Bool := match b with @@ -161,8 +178,8 @@ private def D (x : VarId) (c : CtorInfo) (b : FnBody) : M FnBody := partial def R (e : FnBody) : M FnBody := do match e with | .case tid x xType alts => - let alreadyFound := (← read).casesVars.contains x - withReader (fun ctx => { ctx with casesVars := ctx.casesVars.insert x }) do + let alreadyFound := (← read).alreadyFound.contains x + withReader (fun ctx => { ctx with alreadyFound := ctx.alreadyFound.insert x }) do let alts ← alts.mapM fun alt => do let alt ← alt.mmodifyBody R match alt with @@ -187,16 +204,43 @@ partial def R (e : FnBody) : M FnBody := do let b ← R b return instr.setBody b -end ResetReuse +abbrev N := StateT (PHashSet VarId) Id + +partial def collectResets (e : FnBody) : N Unit := do + match e with + | .case _ _ _ alts => alts.forM fun alt => collectResets alt.body + | .jdecl _ _ v b => collectResets v; collectResets b + | .vdecl _ _ (.reset _ x) b => modify fun s => s.insert x; collectResets b + | e => unless e.isTerminal do + let (_, b) := e.split + collectResets b +end ResetReuse open ResetReuse -def Decl.insertResetReuse (d : Decl) : Decl := + +def Decl.insertResetReuseCore (d : Decl) (relaxedReuse : Bool) : Decl := match d with | .fdecl (body := b) .. => let nextIndex := d.maxIndex + 1 - let bNew := (R b {}).run' nextIndex + -- First time we execute `insertResetReuseCore`, `relaxedReuse := false`. + let alreadyFound : PHashSet VarId := if relaxedReuse then (collectResets b *> get).run' {} else {} + let bNew := R b { relaxedReuse, alreadyFound } |>.run' nextIndex d.updateBody! bNew | other => other +def Decl.insertResetReuse (d : Decl) : Decl := + /- + We execute the reset/reuse algorithm twice. The first time, we only reuse memory cells + between identical constructor memory cells. That is, we do not reuse a `PSigma.mk` memory cell + when allocating a `Prod.mk` memory cell, even though they have the same layout. Recall + that the reset/reuse placement algorithm is a heuristic, and the first pass prevents reuses + that are unlikely to be useful at runtime. Then, we run the procedure again, + relaxing this restriction. If there are still opportunities for reuse, we will take advantage of them. + + The second pass addresses issue #4089. + -/ + d.insertResetReuseCore (relaxedReuse := false) + |>.insertResetReuseCore (relaxedReuse := true) + end Lean.IR diff --git a/tests/lean/4089.lean b/tests/lean/4089.lean new file mode 100644 index 000000000000..f205995323b6 --- /dev/null +++ b/tests/lean/4089.lean @@ -0,0 +1,12 @@ +set_option trace.compiler.ir.reset_reuse true + +def f : Nat × Nat → Nat × Nat + | (a, b) => (b, a) + +def Sigma.toProd : (_ : α) × β → α × β + | ⟨a, b⟩ => (a, b) + +def foo : List (Nat × Nat) → List Nat + | [] => [] + | x :: xs => match x with + | (a, _) => a :: foo xs diff --git a/tests/lean/4089.lean.expected.out b/tests/lean/4089.lean.expected.out new file mode 100644 index 000000000000..9152f7bfa4d2 --- /dev/null +++ b/tests/lean/4089.lean.expected.out @@ -0,0 +1,38 @@ + +[reset_reuse] +def f (x_1 : obj) : obj := + case x_1 : obj of + Prod.mk → + let x_2 : obj := proj[0] x_1; + let x_3 : obj := proj[1] x_1; + let x_5 : obj := reset[2] x_1; + let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_3 x_2; + ret x_4 +[reset_reuse] +def Sigma.toProd._rarg (x_1 : obj) : obj := + case x_1 : obj of + Sigma.mk → + let x_2 : obj := proj[0] x_1; + let x_3 : obj := proj[1] x_1; + let x_5 : obj := reset[2] x_1; + let x_4 : obj := reuse x_5 in ctor_0[Prod.mk] x_2 x_3; + ret x_4 +def Sigma.toProd (x_1 : ◾) (x_2 : ◾) : obj := + let x_3 : obj := pap Sigma.toProd._rarg; + ret x_3 +[reset_reuse] +def foo (x_1 : obj) : obj := + case x_1 : obj of + List.nil → + let x_2 : obj := ctor_0[List.nil]; + ret x_2 + List.cons → + let x_3 : obj := proj[0] x_1; + case x_3 : obj of + Prod.mk → + let x_4 : obj := proj[1] x_1; + let x_9 : obj := reset[2] x_1; + let x_5 : obj := proj[0] x_3; + let x_6 : obj := foo x_4; + let x_7 : obj := reuse x_9 in ctor_1[List.cons] x_5 x_6; + ret x_7 \ No newline at end of file From e5b7dc819bebee0e1dd05fec72b6049bc0d803e9 Mon Sep 17 00:00:00 2001 From: Siddharth Date: Tue, 7 May 2024 23:31:19 +0100 Subject: [PATCH 11/44] feat: bitvec lemma to turn negation into bitwise not+add (#4095) Identity 2-2 (a) (Section: Addition Combined with Logical Operations) from Hacker's Delight, 2nd edition. --- src/Init/Data/BitVec/Lemmas.lean | 7 +++++++ 1 file changed, 7 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 23cc6c991995..27f3632e23d1 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -919,6 +919,13 @@ theorem negOne_eq_allOnes : -1#w = allOnes w := by have r : (2^w - 1) < 2^w := by omega simp [Nat.mod_eq_of_lt q, Nat.mod_eq_of_lt r] +theorem neg_eq_not_add (x : BitVec w) : -x = ~~~x + 1 := by + apply eq_of_toNat_eq + simp only [toNat_neg, ofNat_eq_ofNat, toNat_add, toNat_not, toNat_ofNat, Nat.add_mod_mod] + congr + have hx : x.toNat < 2^w := x.isLt + rw [Nat.sub_sub, Nat.add_comm 1 x.toNat, ← Nat.sub_sub, Nat.sub_add_cancel (by omega)] + /-! ### mul -/ theorem mul_def {n} {x y : BitVec n} : x * y = (ofFin <| x.toFin * y.toFin) := by rfl From ec27b3760d6eaabd2418086a4400228fdf73b465 Mon Sep 17 00:00:00 2001 From: =?UTF-8?q?Fran=C3=A7ois=20G=2E=20Dorais?= Date: Tue, 7 May 2024 19:29:38 -0400 Subject: [PATCH 12/44] fix: swap `Nat.zero_or` and `Nat.or_zero` (#4094) Closes #4093 --- src/Init/Data/Nat/Bitwise/Lemmas.lean | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/src/Init/Data/Nat/Bitwise/Lemmas.lean b/src/Init/Data/Nat/Bitwise/Lemmas.lean index 90610f95df3d..c59bff8945ec 100644 --- a/src/Init/Data/Nat/Bitwise/Lemmas.lean +++ b/src/Init/Data/Nat/Bitwise/Lemmas.lean @@ -402,12 +402,12 @@ theorem and_pow_two_identity {x : Nat} (lt : x < 2^n) : x &&& 2^n-1 = x := by /-! ### lor -/ -@[simp] theorem or_zero (x : Nat) : 0 ||| x = x := by +@[simp] theorem zero_or (x : Nat) : 0 ||| x = x := by simp only [HOr.hOr, OrOp.or, lor] unfold bitwise simp [@eq_comm _ 0] -@[simp] theorem zero_or (x : Nat) : x ||| 0 = x := by +@[simp] theorem or_zero (x : Nat) : x ||| 0 = x := by simp only [HOr.hOr, OrOp.or, lor] unfold bitwise simp [@eq_comm _ 0] From 2a5ca00ad61cbfeb69f150701048e3b67d20e073 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 May 2024 01:31:05 +0200 Subject: [PATCH 13/44] perf: issue at `binop%` and `binrel%` elaborators (#4092) This issue was affecting several Mathlib files. @mattrobball @semorrison This is a different solution for the issue. The comment at `Extra.lean` describes the new solution and documents the new issues found with the previous one. closes #4085 --- src/Lean/Elab/Extra.lean | 28 +- tests/lean/run/binop_binrel_perf_issue.lean | 1015 +++++++++++++++++++ 2 files changed, 1042 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/binop_binrel_perf_issue.lean diff --git a/src/Lean/Elab/Extra.lean b/src/Lean/Elab/Extra.lean index fd03c0f9f0fa..c4acb68e1238 100644 --- a/src/Lean/Elab/Extra.lean +++ b/src/Lean/Elab/Extra.lean @@ -273,7 +273,33 @@ where match (← get).max? with | none => modify fun s => { s with max? := type } | some max => - unless (← withNewMCtxDepth <| isDefEqGuarded max type) do + /- + Remark: Previously, we used `withNewMCtxDepth` to prevent metavariables in `max` and `type` from being assigned. + + Reason: This is a heuristic procedure for introducing coercions in scenarios such as: + - Given `(n : Nat) (i : Int)`, elaborate `n = i`. The coercion must be inserted at `n`. + Consider the elaboration problem `(n + 0) + i`, where the type of term `0` is a metavariable. + We do not want it to be elaborated as `(Int.ofNat n + Int.ofNat (0 : Nat)) + i`; instead, we prefer the result to be `(Int.ofNat n + (0 : Int)) + i`. + Here is another example where we avoid assigning metavariables: `max := BitVec n` and `type := BitVec ?m`. + + However, the combination `withNewMCtxDepth <| isDefEqGuarded max type` introduced performance issues in several + Mathlib files because `isDefEq` was spending a lot of time unfolding definitions in `max` and `type` before failing. + + To address this issue, we allowed only reducible definitions to be unfolded during this check, using + `withNewMCtxDepth <| withReducible <| isDefEqGuarded max type`. This change fixed some performance issues but created new ones. + Lean was now spending time trying to use `hasCoe`, likely occurring in places where `withNewMCtxDepth <| isDefEqGuarded max type` + used to succeed but was now failing after we introduced `withReducible`. + + We then considered using just `isDefEqGuarded max type` and changing the definition of `isUnknown`. In the new definition, + the else-case would be `| e => e.hasExprMVar` instead of `| _ => false`. However, we could not even compile this repo using + this configuration. The problem arises because some files require coercions even when `max` contains metavariables, + for example: `max := Option ?m` and `type := Name`. + + As a result, rather than restricting reducibility, we decided to set `Meta.Config.isDefEqStuckEx := true`. + This means that if `isDefEq` encounters a subproblem `?m =?= a` where `?m` is non-assignable, it aborts the test + instead of unfolding definitions. + -/ + unless (← withNewMCtxDepth <| withConfig (fun config => { config with isDefEqStuckEx := true }) <| isDefEqGuarded max type) do if (← hasCoe type max) then return () else if (← hasCoe max type) then diff --git a/tests/lean/run/binop_binrel_perf_issue.lean b/tests/lean/run/binop_binrel_perf_issue.lean new file mode 100644 index 000000000000..fae2fac9841b --- /dev/null +++ b/tests/lean/run/binop_binrel_perf_issue.lean @@ -0,0 +1,1015 @@ +/-! +This is a minimization of an `isDefEq` timeout from Mathlib.Algebra.Module.Submodule.Localization, +where it is reasonably fast with `set_option backward.isDefEq.lazyWhnfCore false`, +but very slow with the default `isDefEq` algorithm. +-/ + +section Mathlib.Init.Set + +set_option autoImplicit true + +def Set (α : Type u) := α → Prop + +def setOf {α : Type u} (p : α → Prop) : Set α := p + +namespace Set + +protected def Mem (a : α) (s : Set α) : Prop := s a + +instance : Membership α (Set α) := ⟨Set.Mem⟩ + +end Set + +end Mathlib.Init.Set + +section Mathlib.Init.ZeroOne + +set_option autoImplicit true + +class Zero.{u} (α : Type u) where + zero : α + +instance (priority := 300) Zero.toOfNat0 {α} [Zero α] : OfNat α (nat_lit 0) where + ofNat := ‹Zero α›.1 + +class One (α : Type u) where + one : α + +instance (priority := 300) One.toOfNat1 {α} [One α] : OfNat α (nat_lit 1) where + ofNat := ‹One α›.1 + +end Mathlib.Init.ZeroOne + +section Mathlib.Init.Function + +universe u₁ u₂ + +variable {α : Sort u₁} {β : Sort u₂} + +def Function.Injective (f : α → β) : Prop := + ∀ ⦃a₁ a₂⦄, f a₁ = f a₂ → a₁ = a₂ + +def Function.Surjective (f : α → β) : Prop := + ∀ b, ∃ a, f a = b + +end Mathlib.Init.Function + +section Mathlib.Data.Subtype + +variable {α : Sort _} {p : α → Prop} + +protected theorem Subtype.ext : ∀ {a1 a2 : { x // p x }}, (a1 : α) = (a2 : α) → a1 = a2 + | ⟨_, _⟩, ⟨_, _⟩, rfl => rfl + +theorem Subtype.coe_injective : Function.Injective (fun (a : Subtype p) ↦ (a : α)) := fun _ _ ↦ Subtype.ext + +end Mathlib.Data.Subtype + +section Mathlib.Data.Set.Defs + +universe u v w + +namespace Set + +variable {α : Type u} {β : Type v} + +def preimage (f : α → β) (s : Set β) : Set α := setOf fun x => f x ∈ s + +infixl:80 " ⁻¹' " => preimage + +variable {ι : Sort _} {f : ι → α} + +def range (f : ι → α) : Set α := setOf fun x => ∃ y, f y = x + +end Set + +end Mathlib.Data.Set.Defs + +section Mathlib.Data.FunLike.Basic + +class DFunLike (F : Sort _) (α : outParam (Sort _)) (β : outParam <| α → Sort _) where + coe : F → ∀ a : α, β a + +abbrev FunLike F α β := DFunLike F α fun _ => β + +namespace DFunLike + +variable {F α β} [i : DFunLike F α β] + +instance (priority := 100) hasCoeToFun : CoeFun F (fun _ ↦ ∀ a : α, β a) where + coe := @DFunLike.coe _ _ β _ + +end DFunLike + +end Mathlib.Data.FunLike.Basic + +section Mathlib.Logic.Relator + +universe u₁ u₂ v₁ v₂ + +variable {α : Sort u₁} {β : Sort u₂} {γ : Sort v₁} {δ : Sort v₂} +variable (R : α → β → Prop) (S : γ → δ → Prop) + +def LiftFun (f : α → γ) (g : β → δ) : Prop := + ∀⦃a b⦄, R a b → S (f a) (g b) + +infixr:40 " ⇒ " => LiftFun + +end Mathlib.Logic.Relator + +section Mathlib.Data.Quot + +variable {α : Sort _} {β : Sort _} + +namespace Quotient + +variable {γ : Sort _} {φ : Sort _} {s₁ : Setoid α} {s₂ : Setoid β} {s₃ : Setoid γ} + +protected def mk'' (a : α) : Quotient s₁ := + Quot.mk s₁.1 a + +theorem surjective_Quotient_mk'' : Function.Surjective (Quotient.mk'' : α → Quotient s₁) := + Quot.exists_rep + +protected def liftOn' (q : Quotient s₁) (f : α → φ) (h : ∀ a b, @Setoid.r α s₁ a b → f a = f b) : + φ := + Quotient.liftOn q f h + +protected def map' (f : α → β) (h : (s₁.r ⇒ s₂.r) f f) : Quotient s₁ → Quotient s₂ := + (Quot.lift fun x ↦ Quot.mk _ (f x)) fun _ _ h₁ ↦ Quot.sound <| h h₁ + +protected def map₂' (f : α → β → γ) (h : (s₁.r ⇒ s₂.r ⇒ s₃.r) f f) : + Quotient s₁ → Quotient s₂ → Quotient s₃ := + Quotient.lift₂ (fun x y ↦ Quotient.mk _ (f x y)) fun _ _ _ _ h₁ h₂ ↦ Quot.sound <| h h₁ h₂ + +end Quotient + +end Mathlib.Data.Quot + +section Mathlib.Data.SetLike.Basic + +class SetLike (A : Type _) (B : outParam <| Type _) where + protected coe : A → Set B + +namespace SetLike + +variable {A : Type _} {B : Type _} [i : SetLike A B] + +instance : CoeTC A (Set B) where coe := SetLike.coe + +instance (priority := 100) instMembership : Membership B A := + ⟨fun x p => x ∈ (p : Set B)⟩ + +instance (priority := 100) : CoeSort A (Type _) := + ⟨fun p => { x : B // x ∈ p }⟩ + +end SetLike + +end Mathlib.Data.SetLike.Basic + +section Mathlib.Algebra.Group.Defs + +universe u v w + +class HVAdd (α : Type u) (β : Type v) (γ : outParam (Type w)) where + hVAdd : α → β → γ + +class HSMul (α : Type u) (β : Type v) (γ : outParam (Type w)) where + hSMul : α → β → γ + +class VAdd (G : Type u) (P : Type v) where + vadd : G → P → P + +class SMul (M : Type u) (α : Type v) where + smul : M → α → α + +infixl:65 " +ᵥ " => HVAdd.hVAdd +infixr:73 " • " => HSMul.hSMul + +macro_rules | `($x • $y) => `(leftact% HSMul.hSMul $x $y) + +instance instHSMul {α β} [SMul α β] : HSMul α β β where + hSMul := SMul.smul + +instance instHVAdd {α β} [VAdd α β] : HVAdd α β β where + hVAdd := VAdd.vadd + +class AddSemigroup (G : Type u) extends Add G where + protected add_assoc : ∀ a b c : G, a + b + c = a + (b + c) + +class MulOneClass (M : Type u) extends One M, Mul M where + +class AddZeroClass (M : Type u) extends Zero M, Add M where + protected zero_add : ∀ a : M, 0 + a = a + protected add_zero : ∀ a : M, a + 0 = a + + +class AddMonoid (M : Type u) extends AddSemigroup M, AddZeroClass M where + protected nsmul : Nat → M → M + +instance AddMonoid.toNatSMul {M : Type _} [AddMonoid M] : SMul Nat M := + ⟨AddMonoid.nsmul⟩ + +class AddCommMonoid (M : Type u) extends AddMonoid M + +def SubNegMonoid.sub' {G : Type u} [AddMonoid G] [Neg G] (a b : G) : G := a + -b + +class SubNegMonoid (G : Type u) extends AddMonoid G, Neg G, Sub G where + protected sub := SubNegMonoid.sub' + protected zsmul : Int → G → G + +instance SubNegMonoid.SMulInt {M} [SubNegMonoid M] : SMul Int M := + ⟨SubNegMonoid.zsmul⟩ + +class AddGroup (A : Type u) extends SubNegMonoid A where + protected add_left_neg : ∀ a : A, -a + a = 0 + +class AddCommGroup (G : Type u) extends AddGroup G, AddCommMonoid G + +end Mathlib.Algebra.Group.Defs + +section Mathlib.Algebra.Opposites + +variable {α : Type _} + +structure PreOpposite (α : Type _) : Type _ where + op' :: + unop' : α + +def MulOpposite (α : Type _) : Type _ := PreOpposite α +def AddOpposite (α : Type _) : Type _ := PreOpposite α + +postfix:max "ᵐᵒᵖ" => MulOpposite + +postfix:max "ᵃᵒᵖ" => AddOpposite + +namespace MulOpposite + +def unop : αᵐᵒᵖ → α := + PreOpposite.unop' + +end MulOpposite + +namespace AddOpposite + +def op : α → αᵃᵒᵖ := + PreOpposite.op' + +def unop : αᵃᵒᵖ → α := + PreOpposite.unop' + +instance instZero [Zero α] : Zero αᵃᵒᵖ where zero := op 0 +instance instAdd [Add α] : Add αᵃᵒᵖ where add x y := op (unop y + unop x) +instance instNeg [Neg α] : Neg αᵃᵒᵖ where neg x := op <| -(unop x) + +end AddOpposite + +end Mathlib.Algebra.Opposites + +section Mathlib.Algebra.Group.Hom.Defs + +variable {M N A B : Type _} + +structure AddMonoidHom (M : Type _) (N : Type _) [Add M] [Add N] where + protected toFun : M → N + +infixr:25 " →+ " => AddMonoidHom + +structure MonoidHom (M : Type _) (N : Type _) [Mul M] [Mul N] where + protected toFun : M → N + +infixr:25 " →* " => MonoidHom + +instance MonoidHom.instFunLike [Mul M] [Mul N] : FunLike (M →* N) M N where + coe f := f.toFun + +instance AddMonoidHom.instFunLike [Add A] [Add B] : FunLike (A →+ B) A B where + coe f := f.toFun + +end Mathlib.Algebra.Group.Hom.Defs + +section Mathlib.GroupTheory.GroupAction.Defs + +variable {M N α : Type _} + +instance (priority := 910) Mul.toSMul (α : Type _) [Mul α] : SMul α α := + ⟨(· * ·)⟩ + +namespace SMul + +variable [SMul M α] + +def comp.smul (g : N → M) (n : N) (a : α) : α := + g n • a + +variable (α) + +def comp (g : N → M) : SMul N α where smul := SMul.comp.smul g + +end SMul + +namespace VAdd + +variable [VAdd M α] + +def comp.vadd (g : N → M) (n : N) (a : α) : α := + g n +ᵥ a + +end VAdd + +variable [AddMonoid M] [VAdd M α] + +namespace VAdd + +variable (α) + +def compHom [AddMonoid N] (g : N →+ M) : VAdd N α where + vadd := VAdd.comp.vadd g + +end VAdd + +end Mathlib.GroupTheory.GroupAction.Defs + +section Mathlib.GroupTheory.Subsemigroup.Basic + +class MulMemClass (S : Type _) (M : Type _) [Mul M] [SetLike S M] : Prop where + +class AddMemClass (S : Type _) (M : Type _) [Add M] [SetLike S M] : Prop where + +structure Subsemigroup (M : Type _) [Mul M] where + carrier : Set M + +structure AddSubsemigroup (M : Type _) [Add M] where + carrier : Set M + +end Mathlib.GroupTheory.Subsemigroup.Basic + +section Mathlib.Algebra.Group.InjSurj + +namespace Function + +namespace Injective + +variable {M₁ : Type _} {M₂ : Type _} + +protected def mulOneClass [Mul M₁] [One M₁] [MulOneClass M₂] (f : M₁ → M₂) (hf : Injective f) (one : f 1 = 1) + (mul : ∀ x y, f (x * y) = f x * f y) : MulOneClass M₁ := + { ‹One M₁›, ‹Mul M₁› with } + +variable [Add M₁] [Zero M₁] + +protected def addZeroClass [AddZeroClass M₂] (f : M₁ → M₂) (hf : Injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) : AddZeroClass M₁ := + { ‹Zero M₁›, ‹Add M₁› with + zero_add := sorry, + add_zero := sorry } + +protected def addSemigroup [AddSemigroup M₂] (f : M₁ → M₂) (hf : Injective f) + (add : ∀ x y, f (x + y) = f x + f y) : AddSemigroup M₁ := + { ‹Add M₁› with add_assoc := sorry } + +variable [SMul Nat M₁] + +protected def addMonoid [AddMonoid M₂] (f : M₁ → M₂) (hf : Injective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (nsmul : ∀ (x) (n : Nat), f (n • x) = n • f x) : AddMonoid M₁ := + { hf.addZeroClass f zero add, hf.addSemigroup f add with + nsmul := fun n x => n • x } + +end Injective + +namespace Surjective + +variable {M₁ : Type _} {M₂ : Type _} [Add M₂] + +protected def addSemigroup [AddSemigroup M₁] (f : M₁ → M₂) (hf : Surjective f) + (add : ∀ x y, f (x + y) = f x + f y) : AddSemigroup M₂ := + { ‹Add M₂› with add_assoc := sorry } + +variable [Zero M₂] + +protected def addZeroClass [AddZeroClass M₁] (f : M₁ → M₂) (hf : Surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) : AddZeroClass M₂ := + { ‹Zero M₂›, ‹Add M₂› with + zero_add := sorry, + add_zero := sorry } + +variable [SMul Nat M₂] + +protected def addMonoid [AddMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) (nsmul : ∀ (x) (n : Nat), f (n • x) = n • f x) : AddMonoid M₂ := + { hf.addSemigroup f add, hf.addZeroClass f zero add with + nsmul := fun n x => n • x } + +variable [Neg M₂] [Sub M₂] [SMul Int M₂] + +protected def subNegMonoid [SubNegMonoid M₁] (f : M₁ → M₂) (hf : Surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) + (nsmul : ∀ (x) (n : Nat), f (n • x) = n • f x) : SubNegMonoid M₂ := + { hf.addMonoid f zero add nsmul, ‹Sub M₂›, ‹Neg M₂› with + zsmul := fun n x => n • x } + +protected def addGroup [AddGroup M₁] (f : M₁ → M₂) (hf : Surjective f) (zero : f 0 = 0) + (add : ∀ x y, f (x + y) = f x + f y) + (nsmul : ∀ (x) (n : Nat), f (n • x) = n • f x) : AddGroup M₂ := + { hf.subNegMonoid f zero add nsmul with + add_left_neg := sorry } + +end Surjective + +end Function + + +end Mathlib.Algebra.Group.InjSurj + +section Mathlib.Algebra.Group.Opposite + +variable {α : Type _} + +namespace AddOpposite + +instance instAddSemigroup [AddSemigroup α] : AddSemigroup αᵃᵒᵖ where + add_assoc x y z := sorry + +instance instAddZeroClass [AddZeroClass α] : AddZeroClass αᵃᵒᵖ where + toAdd := instAdd + toZero := instZero + zero_add _ := sorry + add_zero _ := sorry + +instance instAddMonoid [AddMonoid α] : AddMonoid αᵃᵒᵖ := + { instAddZeroClass with + toAddSemigroup := instAddSemigroup + nsmul := fun n a => op <| n • a.unop } + +instance instSubNegMonoid [SubNegMonoid α] : SubNegMonoid αᵃᵒᵖ where + toAddMonoid := instAddMonoid + toNeg := instNeg + zsmul n a := op <| n • a.unop + +instance instAddGroup [AddGroup α] : AddGroup αᵃᵒᵖ where + toSubNegMonoid := instSubNegMonoid + add_left_neg _ := sorry + +end AddOpposite + +end Mathlib.Algebra.Group.Opposite + +section Mathlib.GroupTheory.Subsemigroup.Operations + +variable {M : Type _} + +namespace MulMemClass + +variable {A : Type _} [Mul M] [SetLike A M] [hA : MulMemClass A M] (S' : A) + +instance (priority := 900) mul : Mul S' := + ⟨fun a b => ⟨a.1 * b.1, sorry⟩⟩ + +end MulMemClass + +namespace AddMemClass + +variable {A : Type _} [Add M] [SetLike A M] [hA : AddMemClass A M] (S' : A) + +instance (priority := 900) add : Add S' := + ⟨fun a b => ⟨a.1 + b.1, sorry⟩⟩ + +end AddMemClass + +end Mathlib.GroupTheory.Subsemigroup.Operations + +section Mathlib.GroupTheory.Submonoid.Basic + +variable {M : Type _} {N : Type _} +variable {A : Type _} + +variable [MulOneClass M] {s : Set M} +variable [AddZeroClass A] {t : Set A} + +structure Submonoid (M : Type _) [MulOneClass M] extends Subsemigroup M where + +class SubmonoidClass (S : Type _) (M : Type _) [MulOneClass M] [SetLike S M] extends + MulMemClass S M : Prop + +structure AddSubmonoid (M : Type _) [AddZeroClass M] extends AddSubsemigroup M where + +class AddSubmonoidClass (S : Type _) (M : Type _) [AddZeroClass M] [SetLike S M] extends + AddMemClass S M : Prop + +namespace AddSubmonoid + +instance : SetLike (AddSubmonoid A) A where + coe s := s.carrier + +instance : AddSubmonoidClass (AddSubmonoid A) A where + +end AddSubmonoid + +namespace Submonoid + +instance : SetLike (Submonoid M) M where + coe s := s.carrier + +instance : SubmonoidClass (Submonoid M) M where + +end Submonoid + +end Mathlib.GroupTheory.Submonoid.Basic + +section Mathlib.GroupTheory.Submonoid.Operations + +variable {M : Type _} [MulOneClass M] (S : Submonoid M) +variable {B : Type _} [AddZeroClass B] (T : AddSubmonoid B) + +section + +variable {A M₁ : Type _} [SetLike A M₁] (S' : A) + +instance one [One M₁] : One S' := + ⟨⟨1, sorry⟩⟩ + +instance zero [Zero M₁] : Zero S' := + ⟨⟨0, sorry⟩⟩ + +end + +namespace AddSubmonoidClass + +instance nSMul {M} [AddMonoid M] {A : Type _} [SetLike A M] + [AddSubmonoidClass A M] (S : A) : SMul Nat S := + ⟨fun n a => ⟨n • a.1, sorry⟩⟩ + +end AddSubmonoidClass + +namespace SubmonoidClass + +instance (priority := 75) toMulOneClass {M : Type _} [MulOneClass M] {A : Type _} [SetLike A M] + [SubmonoidClass A M] (S : A) : MulOneClass S := + Subtype.coe_injective.mulOneClass Subtype.val sorry sorry + +instance (priority := 75) toAddMonoid {M : Type _} [AddMonoid M] {A : Type _} [SetLike A M] + [AddSubmonoidClass A M] (S : A) : AddMonoid S := + Subtype.coe_injective.addMonoid Subtype.val sorry sorry sorry + +end SubmonoidClass + +namespace AddSubmonoid + +def subtype : T →+ B where + toFun := Subtype.val + +end AddSubmonoid + +namespace Submonoid + +def subtype : S →* M where + toFun := Subtype.val + +end Submonoid + +namespace Submonoid + +variable {M' : Type _} {α : Type _} + +section MulOneClass + +variable [MulOneClass M'] + +instance smul [SMul M' α] (S : Submonoid M') : SMul S α := + SMul.comp _ S.subtype + +end MulOneClass + +variable [AddMonoid M'] + +instance vadd [VAdd M' α] (S : AddSubmonoid M') : VAdd S α := + VAdd.compHom _ S.subtype + +end Submonoid + +end Mathlib.GroupTheory.Submonoid.Operations + +section Mathlib.GroupTheory.Congruence + +variable (M : Type _) {N : Type _} {P : Type _} + +structure AddCon [Add M] extends Setoid M where + +variable {M} + +namespace AddCon + +section + +variable [Add M] (c : AddCon M) + +protected def Quotient := + Quotient c.toSetoid + +protected def liftOn {β} {c : AddCon M} (q : c.Quotient) (f : M → β) (h : ∀ a b, c.r a b → f a = f b) : + β := + Quotient.liftOn' q f h + +instance hasAdd : Add c.Quotient := + ⟨Quotient.map₂' (· + ·) sorry⟩ + +end + +section AddZeroClass + +variable [AddZeroClass M] [AddZeroClass P] (c : AddCon M) + +variable (f : M →+ P) + +def lift : c.Quotient →+ P where + toFun x := (AddCon.liftOn x f) sorry + +end AddZeroClass + +section Monoids + +instance zero [AddZeroClass M] (c : AddCon M) : Zero c.Quotient where + zero := Quotient.mk'' (0 : M) + +instance _root_.AddCon.Quotient.nsmul {M : Type _} [AddMonoid M] (c : AddCon M) : + SMul Nat c.Quotient where + smul n := (Quotient.map' (n • ·)) sorry + +instance addSemigroup {M : Type _} [AddSemigroup M] (c : AddCon M) : AddSemigroup c.Quotient := + { (Function.Surjective.addSemigroup _ Quotient.surjective_Quotient_mk'' sorry : + AddSemigroup c.Quotient) with + toAdd := AddCon.hasAdd _ } + +instance addMonoid {M : Type _} [AddMonoid M] (c : AddCon M) : AddMonoid c.Quotient := + { (Function.Surjective.addMonoid _ Quotient.surjective_Quotient_mk'' sorry sorry sorry : AddMonoid c.Quotient) with + toAddSemigroup := AddCon.addSemigroup _ + toZero := AddCon.zero _ } + +end Monoids + +variable [AddGroup M] [AddGroup N] [AddGroup P] (c : AddCon M) + +instance hasNeg : Neg c.Quotient := + ⟨(Quotient.map' Neg.neg) sorry⟩ + +instance hasSub : Sub c.Quotient := + ⟨(Quotient.map₂' (· - ·)) sorry⟩ + +instance _root_.AddCon.Quotient.zsmul {M : Type _} [AddGroup M] (c : AddCon M) : + SMul Int c.Quotient := + ⟨fun z => (Quotient.map' (z • ·)) sorry⟩ + +instance addGroup : AddGroup c.Quotient := + { (Function.Surjective.addGroup Quotient.mk'' + Quotient.surjective_Quotient_mk'' sorry sorry sorry : AddGroup c.Quotient) with + toAddMonoid := AddCon.addMonoid _ + toNeg := AddCon.hasNeg _ + toSub := AddCon.hasSub _ } + +end AddCon + +end Mathlib.GroupTheory.Congruence + +section Mathlib.GroupTheory.GroupAction.Opposite + +variable {α} + +instance Add.toHasOppositeVAdd [Add α] : VAdd αᵃᵒᵖ α := + ⟨fun c x => x + c.unop⟩ + +instance AddMonoid.toOppositeVAdd [AddMonoid α] : VAdd αᵃᵒᵖ α where + vadd := (· +ᵥ ·) + +end Mathlib.GroupTheory.GroupAction.Opposite + +section Mathlib.GroupTheory.Subgroup.Basic + +variable {A : Type _} [AddGroup A] + +structure AddSubgroup (G : Type _) [AddGroup G] extends AddSubmonoid G where + +namespace AddSubgroup + +instance : SetLike (AddSubgroup A) A where + coe s := s.carrier + +end AddSubgroup + +end Mathlib.GroupTheory.Subgroup.Basic + +section Mathlib.GroupTheory.GroupAction.Basic + +universe u v + +namespace VAdd + +variable (M : Type u) [AddMonoid M] {α : Type v} [VAdd M α] + +def orbit (a : α) := + Set.range fun m : M => m +ᵥ a + +variable (G α : Type _) [AddGroup G] [VAdd G α] + +def orbitRel : Setoid α where + r a b := a ∈ orbit G b + iseqv := sorry + +end VAdd + +end Mathlib.GroupTheory.GroupAction.Basic + +section Mathlib.GroupTheory.Subgroup.MulOpposite + +variable {G : Type _} [AddGroup G] + +protected def AddSubgroup.op (H : AddSubgroup G) : AddSubgroup Gᵃᵒᵖ where + carrier := MulOpposite.unop ⁻¹' (H : Set G) + +end Mathlib.GroupTheory.Subgroup.MulOpposite + +section Mathlib.Algebra.Quotient + +universe u v + +class HasQuotient (A : outParam <| Type u) (B : Type v) where + quotient' : B → Type max u v + +def HasQuotient.Quotient (A : outParam <| Type u) {B : Type v} + [HasQuotient A B] (b : B) : Type max u v := + HasQuotient.quotient' b + +notation:35 G " ⧸ " H:34 => HasQuotient.Quotient G H + +end Mathlib.Algebra.Quotient + +section Mathlib.GroupTheory.Coset + +variable {α : Type _} + +namespace QuotientAddGroup + +variable [AddGroup α] (s : AddSubgroup α) + +def leftRel : Setoid α := + VAdd.orbitRel s.op α + +instance instHasQuotientAddSubgroup : HasQuotient α (AddSubgroup α) := + ⟨fun s => Quotient (leftRel s)⟩ + +end QuotientAddGroup + +end Mathlib.GroupTheory.Coset + +section Mathlib.Algebra.Module.LinearMap.Basic + +open Function + +universe u u' v w x y z + +variable {R R₁ R₂ R₃ k S M M₁ M₂ M₃ : Type _} + +structure LinearMap (R : Type _) [Mul R] (M : Type _) + (M₂ : Type _) [AddCommMonoid M] [AddCommMonoid M₂] [SMul R M] [SMul R M₂] extends + AddMonoidHom M M₂ + +notation:25 M " →ₗ[" R:25 "] " M₂:0 => LinearMap R M M₂ + +namespace LinearMap + +section AddCommMonoid + +variable [Mul R] [Mul S] + +section + +variable [AddCommMonoid M] [AddCommMonoid M₃] +variable [SMul R M] [SMul R M₃] + +instance instFunLike : FunLike (M →ₗ[R] M₃) M M₃ where + coe f := f.toFun + +end + +section + +variable [AddCommMonoid M] [AddCommMonoid M₁] [AddCommMonoid M₂] [AddCommMonoid M₃] +variable {module_M₁ : SMul R M₁} {module_M₂ : SMul R M₂} {module_M₃ : SMul R M₃} +variable (f : M₂ →ₗ[R] M₃) (g : M₁ →ₗ[R] M₂) + +def comp : M₁ →ₗ[R] M₃ where + toFun := f ∘ g + +end + +end AddCommMonoid + +end LinearMap + +end Mathlib.Algebra.Module.LinearMap.Basic + +section Mathlib.Algebra.Module.Submodule.Basic + +universe u v + +variable {R : Type u} {M : Type v} + +structure Submodule (R : Type u) (M : Type v) [Mul R] [AddCommMonoid M] [SMul R M] extends + AddSubmonoid M : Type v + +instance setLike [Mul R] [AddCommMonoid M] [SMul R M] : SetLike (Submodule R M) M where + coe s := s.carrier + +def Submodule.toAddSubgroup [Mul R] [AddCommGroup M] {module_M : SMul R M} (p : Submodule R M) : AddSubgroup M := + { p.toAddSubmonoid with } + +end Mathlib.Algebra.Module.Submodule.Basic + +section Mathlib.Algebra.Module.Submodule.RestrictScalars + +variable (S : Type _) {R M : Type _} [Mul R] [AddCommMonoid M] [Mul S] + [SMul S M] [SMul R M] + +def Submodule.restrictScalars (V : Submodule R M) : Submodule S M where + carrier := V + +end Mathlib.Algebra.Module.Submodule.RestrictScalars + +section Mathlib.GroupTheory.QuotientGroup + +universe u x + +namespace QuotientAddGroup + +variable {G : Type u} [AddGroup G] (N : AddSubgroup G) {M : Type x} [AddMonoid M] + +protected def con : AddCon G where + toSetoid := leftRel N + +instance Quotient.addGroup : AddGroup (G ⧸ N) := + (QuotientAddGroup.con N).addGroup + +instance Quotient.addCommGroup {G : Type _} [AddCommGroup G] (N : AddSubgroup G) : AddCommGroup (G ⧸ N) := + { toAddGroup := @QuotientAddGroup.Quotient.addGroup _ _ N } + +def lift (φ : G →+ M) : G ⧸ N →+ M := + (QuotientAddGroup.con N).lift φ + +end QuotientAddGroup + +end Mathlib.GroupTheory.QuotientGroup + +section Mathlib.Algebra.Algebra.Basic + +universe u v + +class Algebra (R : Type u) (A : Type v) [Mul R] [Mul A] extends SMul R A, + R →* A where + +end Mathlib.Algebra.Algebra.Basic + +section Mathlib.Algebra.Module.LocalizedModule + +namespace LocalizedModule + +universe u v + +variable {R : Type u} [MulOneClass R] (S : Submonoid R) +variable (M : Type v) [AddCommMonoid M] [SMul R M] + +def r (a b : M × S) : Prop := + ∃ u : S, u • b.2 • a.1 = u • a.2 • b.1 + +instance r.setoid : Setoid (M × S) where + r := r S M + iseqv := sorry + +def _root_.LocalizedModule : Type max u v := + Quotient (r.setoid S M) + +section + +variable {M S} + +def mk (m : M) (s : S) : LocalizedModule S M := + Quotient.mk' ⟨m, s⟩ + +instance : AddCommMonoid (LocalizedModule S M) := sorry + +noncomputable instance isModule' : SMul R (LocalizedModule S M) := sorry + +end + +end LocalizedModule + +universe u v + +variable {R : Type _} [MulOneClass R] (S : Submonoid R) +variable {M M' M'' : Type _} [AddCommMonoid M] [AddCommMonoid M'] +variable [SMul R M] [SMul R M'] +variable (f : M →ₗ[R] M') + +class IsLocalizedModule (S : Submonoid R) (f : M →ₗ[R] M') : Prop where + +namespace IsLocalizedModule + +variable [IsLocalizedModule S f] + +noncomputable def fromLocalizedModule (f : M →ₗ[R] M') [IsLocalizedModule S f] : + LocalizedModule S M →ₗ[R] M' := + sorry + +variable {S} + +noncomputable def mk' (m : M) (s : S) : M' := + fromLocalizedModule S f (LocalizedModule.mk m s) + +end IsLocalizedModule + +end Mathlib.Algebra.Module.LocalizedModule + +section Mathlib.LinearAlgebra.Quotient + +namespace Submodule + +variable {R M : Type _} {r : R} {x y : M} [Mul R] [AddCommGroup M] [SMul R M] +variable (p : Submodule R M) + +open LinearMap + +def quotientRel : Setoid M := + QuotientAddGroup.leftRel p.toAddSubgroup + +instance hasQuotient : HasQuotient M (Submodule R M) := + ⟨fun p => Quotient (quotientRel p)⟩ + +namespace Quotient + +def mk {p : Submodule R M} : M → M ⧸ p := + Quotient.mk'' + +instance addCommGroup : AddCommGroup (M ⧸ p) := + QuotientAddGroup.Quotient.addCommGroup p.toAddSubgroup + +section Module + +variable {S : Type _} + +instance module' [Mul S] [SMul S R] [SMul S M] (P : Submodule R M) : + SMul S (M ⧸ P) := sorry + +instance module (P : Submodule R M) : SMul R (M ⧸ P) := + Quotient.module' P + +end Module + +end Quotient + +section + +variable {M₂ : Type _} [AddCommGroup M₂] [SMul R M₂] + +def mkQ : M →ₗ[R] M ⧸ p where + toFun := Quotient.mk + +end + +variable {M₂ : Type _} [AddCommGroup M₂] [SMul R M₂] + +def liftQ (f : M →ₗ[R] M₂) : M ⧸ p →ₗ[R] M₂ := + { QuotientAddGroup.lift p.toAddSubgroup f.toAddMonoidHom with } + +variable (q : Submodule R M₂) + +def mapQ (f : M →ₗ[R] M₂) : M ⧸ p →ₗ[R] M₂ ⧸ q := + p.liftQ (q.mkQ.comp f) + +end Submodule + +end Mathlib.LinearAlgebra.Quotient + +section Mathlib.Algebra.Module.Submodule.Localization + +universe u u' v v' + +variable {R : Type u} (S : Type u') {M : Type v} {N : Type v'} +variable [MulOneClass R] [Mul S] [AddCommGroup M] [AddCommGroup N] +variable [SMul R M] [SMul R N] [Algebra R S] [SMul S N] +variable (p : Submonoid R) (f : M →ₗ[R] N) [IsLocalizedModule p f] +variable (M' : Submodule R M) + +/-- Let `S` be the localization of `R` at `p` and `N` be the localization of `M` at `p`. +This is the localization of an `R`-submodule of `M` viewed as an `S`-submodule of `N`. -/ +def Submodule.localized' : Submodule S N where + carrier := setOf fun x => ∃ m ∈ M', ∃ s : p, IsLocalizedModule.mk' f m s = x + +/-- The localization map of a quotient module. -/ +def Submodule.toLocalizedQuotient' : M ⧸ M' →ₗ[R] N ⧸ M'.localized' S p f := + Submodule.mapQ M' ((M'.localized' S p f).restrictScalars R) f + +-- Should after `binrel%` and `binop%` were fixed +set_option maxHeartbeats 1000 in +theorem Submodule.toLocalizedQuotient'_mk₄ (x : M) : + (M'.toLocalizedQuotient' S p f (Submodule.Quotient.mk x)) = (Submodule.Quotient.mk (f x)) := + rfl + +end Mathlib.Algebra.Module.Submodule.Localization From b8e67d87a8b41fa20dc629f9f63d9aff1855d2b2 Mon Sep 17 00:00:00 2001 From: Austin Letson Date: Tue, 7 May 2024 19:49:43 -0400 Subject: [PATCH 14/44] doc: add docstrings and usage examples in `Init.Data.String.Basic` (#4001) Add docstrings and usage examples for `String.length`, `.push`, `.append`, `.get?`, `.set`, `.modyify`, and `.next`. Update docstrings and add usage examples for `String.toList`, `.get`, and `.get!`. --------- Co-authored-by: Joachim Breitner Co-authored-by: David Thrane Christiansen --- src/Init/Data/String/Basic.lean | 107 +++++++++++++++++++++++++++++--- 1 file changed, 98 insertions(+), 9 deletions(-) diff --git a/src/Init/Data/String/Basic.lean b/src/Init/Data/String/Basic.lean index 1dbe48e753a1..1ea50f2e6869 100644 --- a/src/Init/Data/String/Basic.lean +++ b/src/Init/Data/String/Basic.lean @@ -24,23 +24,51 @@ instance : LT String := instance decLt (s₁ s₂ : @& String) : Decidable (s₁ < s₂) := List.hasDecidableLt s₁.data s₂.data +/-- +Returns the length of a string in Unicode code points. + +Examples: +* `"".length = 0` +* `"abc".length = 3` +* `"L∃∀N".length = 4` +-/ @[extern "lean_string_length"] def length : (@& String) → Nat | ⟨s⟩ => s.length -/-- The internal implementation uses dynamic arrays and will perform destructive updates - if the String is not shared. -/ +/-- +Pushes a character onto the end of a string. + +The internal implementation uses dynamic arrays and will perform destructive updates +if the string is not shared. + +Example: `"abc".push 'd' = "abcd"` +-/ @[extern "lean_string_push"] def push : String → Char → String | ⟨s⟩, c => ⟨s ++ [c]⟩ -/-- The internal implementation uses dynamic arrays and will perform destructive updates - if the String is not shared. -/ +/-- +Appends two strings. + +The internal implementation uses dynamic arrays and will perform destructive updates +if the string is not shared. + +Example: `"abc".append "def" = "abcdef"` +-/ @[extern "lean_string_append"] def append : String → (@& String) → String | ⟨a⟩, ⟨b⟩ => ⟨a ++ b⟩ -/-- O(n) in the runtime, where n is the length of the String -/ +/-- +Converts a string to a list of characters. + +Even though the logical model of strings is as a structure that wraps a list of characters, +this operation takes time and space linear in the length of the string, because the compiler +uses an optimized representation as dynamic arrays. + +Example: `"abc".toList = ['a', 'b', 'c']` +-/ def toList (s : String) : List Char := s.data @@ -59,9 +87,17 @@ def utf8GetAux : List Char → Pos → Pos → Char | c::cs, i, p => if i = p then c else utf8GetAux cs (i + c) p /-- - Return character at position `p`. If `p` is not a valid position - returns `(default : Char)`. - See `utf8GetAux` for the reference implementation. +Returns the character at position `p` of a string. If `p` is not a valid position, +returns `(default : Char)`. + +See `utf8GetAux` for the reference implementation. + +Examples: +* `"abc".get ⟨1⟩ = 'b'` +* `"abc".get ⟨3⟩ = (default : Char) = 'A'` + +Positions can also be invalid if a byte index points into the middle of a multi-byte UTF-8 +character. For example,`"L∃∀N".get ⟨2⟩ = (default : Char) = 'A'`. -/ @[extern "lean_string_utf8_get"] def get (s : @& String) (p : @& Pos) : Char := @@ -72,12 +108,30 @@ def utf8GetAux? : List Char → Pos → Pos → Option Char | [], _, _ => none | c::cs, i, p => if i = p then c else utf8GetAux? cs (i + c) p +/-- +Returns the character at position `p`. If `p` is not a valid position, returns `none`. + +Examples: +* `"abc".get? ⟨1⟩ = some 'b'` +* `"abc".get? ⟨3⟩ = none` + +Positions can also be invalid if a byte index points into the middle of a multi-byte UTF-8 +character. For example, `"L∃∀N".get? ⟨2⟩ = none` +-/ @[extern "lean_string_utf8_get_opt"] def get? : (@& String) → (@& Pos) → Option Char | ⟨s⟩, p => utf8GetAux? s 0 p /-- - Similar to `get`, but produces a panic error message if `p` is not a valid `String.Pos`. +Returns the character at position `p` of a string. If `p` is not a valid position, +returns `(default : Char)` and produces a panic error message. + +Examples: +* `"abc".get! ⟨1⟩ = 'b'` +* `"abc".get! ⟨3⟩` panics + +Positions can also be invalid if a byte index points into the middle of a multi-byte UTF-8 character. For example, +`"L∃∀N".get! ⟨2⟩` panics. -/ @[extern "lean_string_utf8_get_bang"] def get! (s : @& String) (p : @& Pos) : Char := @@ -89,13 +143,48 @@ def utf8SetAux (c' : Char) : List Char → Pos → Pos → List Char | c::cs, i, p => if i = p then (c'::cs) else c::(utf8SetAux c' cs (i + c) p) +/-- +Replaces the character at a specified position in a string with a new character. If the position +is invalid, the string is returned unchanged. + +If both the replacement character and the replaced character are ASCII characters and the string +is not shared, destructive updates are used. + +Examples: +* `"abc".set ⟨1⟩ 'B' = "aBc"` +* `"abc".set ⟨3⟩ 'D' = "abc"` +* `"L∃∀N".set ⟨4⟩ 'X' = "L∃XN"` + +Because `'∃'` is a multi-byte character, the byte index `2` in `L∃∀N` is an invalid position, +so `"L∃∀N".set ⟨2⟩ 'X' = "L∃∀N"`. +-/ @[extern "lean_string_utf8_set"] def set : String → (@& Pos) → Char → String | ⟨s⟩, i, c => ⟨utf8SetAux c s 0 i⟩ +/-- +Replaces the character at position `p` in the string `s` with the result of applying `f` to that character. +If `p` is an invalid position, the string is returned unchanged. + +Examples: +* `abc.modify ⟨1⟩ Char.toUpper = "aBc"` +* `abc.modify ⟨3⟩ Char.toUpper = "abc"` +-/ def modify (s : String) (i : Pos) (f : Char → Char) : String := s.set i <| f <| s.get i +/-- +Returns the next position in a string after position `p`. If `p` is not a valid position or `p = s.endPos`, +the result is unspecified. + +Examples: +* `"abc".next ⟨1⟩ = String.Pos.mk 2` +* `"L∃∀N".next ⟨1⟩ = String.Pos.mk 4`, since `'∃'` is a multi-byte UTF-8 character + +Cases where the result is unspecified: +* `"abc".next ⟨3⟩`, since `3 = s.endPos` +* `"L∃∀N".next ⟨2⟩`, since `2` points into the middle of a multi-byte UTF-8 character +-/ @[extern "lean_string_utf8_next"] def next (s : @& String) (p : @& Pos) : Pos := let c := get s p From a257767417918c63d356bbf6db7cfbb0b8c74fb4 Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Tue, 7 May 2024 16:55:52 -0700 Subject: [PATCH 15/44] fix: make deriving handler for `Repr` be consistent about erasing types and proofs (#3944) The deriving handler would use `_` for types and proofs for structures but not for inductives. Reported by Graham Leach-Krouse on [Zulip](https://leanprover.zulipchat.com/#narrow/stream/113489-new-members/topic/Deriving.20Repr.20for.20an.20inductive.20with.20proof.20parameters/near/434181985). --- src/Lean/Elab/Deriving/Repr.lean | 2 + tests/lean/derivingRepr.lean | 39 ---------- tests/lean/derivingRepr.lean.expected.out | 23 ------ tests/lean/run/derivingRepr.lean | 90 +++++++++++++++++++++++ 4 files changed, 92 insertions(+), 62 deletions(-) delete mode 100644 tests/lean/derivingRepr.lean delete mode 100644 tests/lean/derivingRepr.lean.expected.out create mode 100644 tests/lean/run/derivingRepr.lean diff --git a/src/Lean/Elab/Deriving/Repr.lean b/src/Lean/Elab/Deriving/Repr.lean index 425d8772ef9e..0a03ed1cc0e3 100644 --- a/src/Lean/Elab/Deriving/Repr.lean +++ b/src/Lean/Elab/Deriving/Repr.lean @@ -70,6 +70,8 @@ where if localDecl.binderInfo.isExplicit then if (← inferType x).isAppOf indVal.name then rhs ← `($rhs ++ Format.line ++ $(mkIdent auxFunName):ident $a:ident max_prec) + else if (← isType x <||> isProof x) then + rhs ← `($rhs ++ Format.line ++ "_") else rhs ← `($rhs ++ Format.line ++ reprArg $a) patterns := patterns.push (← `(@$(mkIdent ctorName):ident $ctorArgs:term*)) diff --git a/tests/lean/derivingRepr.lean b/tests/lean/derivingRepr.lean deleted file mode 100644 index 59a83beffa6a..000000000000 --- a/tests/lean/derivingRepr.lean +++ /dev/null @@ -1,39 +0,0 @@ -structure Foo where - name : String - val : List Nat - lower : Nat := List.length val - inv : val.length >= lower - flag : Bool - deriving Repr - -#eval { name := "Joe", val := List.iota 40, flag := true, inv := by decide : Foo } - -inductive Tree (α : Type) where - | node : List (Tree α) → Bool → Tree α - | leaf : α → Tree α - deriving Repr - -#eval Tree.node (List.iota 10 |>.map fun i => Tree.node [Tree.leaf i] (i%2==0)) true - -inductive StructureLikeInductive where - | field : Nat -> StructureLikeInductive - deriving Repr - -#eval StructureLikeInductive.field 5 - -namespace Foo -mutual -inductive Tree (α : Type u) where - | node : TreeList α → Tree α - | leaf : α → Tree α - deriving Repr - -inductive TreeList (α : Type u) where - | nil : TreeList α - | cons : Tree α → TreeList α → TreeList α - deriving Repr -end - -#eval Tree.node (TreeList.cons (Tree.leaf 30) (TreeList.cons (Tree.leaf 20) (TreeList.cons (Tree.leaf 10) TreeList.nil))) - -end Foo diff --git a/tests/lean/derivingRepr.lean.expected.out b/tests/lean/derivingRepr.lean.expected.out deleted file mode 100644 index f5bdbc24294b..000000000000 --- a/tests/lean/derivingRepr.lean.expected.out +++ /dev/null @@ -1,23 +0,0 @@ -{ name := "Joe", - val := [40, 39, 38, 37, 36, 35, 34, 33, 32, 31, 30, 29, 28, 27, 26, 25, 24, 23, 22, 21, 20, 19, 18, 17, 16, 15, 14, - 13, 12, 11, 10, 9, 8, 7, 6, 5, 4, 3, 2, 1], - lower := 40, - inv := _, - flag := true } -Tree.node - [Tree.node [Tree.leaf 10] true, - Tree.node [Tree.leaf 9] false, - Tree.node [Tree.leaf 8] true, - Tree.node [Tree.leaf 7] false, - Tree.node [Tree.leaf 6] true, - Tree.node [Tree.leaf 5] false, - Tree.node [Tree.leaf 4] true, - Tree.node [Tree.leaf 3] false, - Tree.node [Tree.leaf 2] true, - Tree.node [Tree.leaf 1] false] - true -StructureLikeInductive.field 5 -Foo.Tree.node - (Foo.TreeList.cons - (Foo.Tree.leaf 30) - (Foo.TreeList.cons (Foo.Tree.leaf 20) (Foo.TreeList.cons (Foo.Tree.leaf 10) (Foo.TreeList.nil)))) diff --git a/tests/lean/run/derivingRepr.lean b/tests/lean/run/derivingRepr.lean new file mode 100644 index 000000000000..37ef4c2b562e --- /dev/null +++ b/tests/lean/run/derivingRepr.lean @@ -0,0 +1,90 @@ +structure Foo where + name : String + val : List Nat + lower : Nat := List.length val + inv : val.length >= lower + flag : Bool + deriving Repr + +/-- +info: { name := "Joe", + val := [40, 39, 38, 37, 36, 35, 34, 33, 32, 31, 30, 29, 28, 27, 26, 25, 24, 23, 22, 21, 20, 19, 18, 17, 16, 15, 14, + 13, 12, 11, 10, 9, 8, 7, 6, 5, 4, 3, 2, 1], + lower := 40, + inv := _, + flag := true } +-/ +#guard_msgs in +#eval { name := "Joe", val := List.iota 40, flag := true, inv := by decide : Foo } + +inductive Tree (α : Type) where + | node : List (Tree α) → Bool → Tree α + | leaf : α → Tree α + deriving Repr + +/-- +info: Tree.node + [Tree.node [Tree.leaf 10] true, + Tree.node [Tree.leaf 9] false, + Tree.node [Tree.leaf 8] true, + Tree.node [Tree.leaf 7] false, + Tree.node [Tree.leaf 6] true, + Tree.node [Tree.leaf 5] false, + Tree.node [Tree.leaf 4] true, + Tree.node [Tree.leaf 3] false, + Tree.node [Tree.leaf 2] true, + Tree.node [Tree.leaf 1] false] + true +-/ +#guard_msgs in +#eval Tree.node (List.iota 10 |>.map fun i => Tree.node [Tree.leaf i] (i%2==0)) true + +inductive StructureLikeInductive where + | field : Nat -> StructureLikeInductive + deriving Repr + +/-- info: StructureLikeInductive.field 5 -/ +#guard_msgs in +#eval StructureLikeInductive.field 5 + +namespace Foo +mutual +inductive Tree (α : Type u) where + | node : TreeList α → Tree α + | leaf : α → Tree α + deriving Repr + +inductive TreeList (α : Type u) where + | nil : TreeList α + | cons : Tree α → TreeList α → TreeList α + deriving Repr +end + +/-- +info: Foo.Tree.node + (Foo.TreeList.cons + (Foo.Tree.leaf 30) + (Foo.TreeList.cons (Foo.Tree.leaf 20) (Foo.TreeList.cons (Foo.Tree.leaf 10) (Foo.TreeList.nil)))) +-/ +#guard_msgs in +#eval Tree.node (TreeList.cons (Tree.leaf 30) (TreeList.cons (Tree.leaf 20) (TreeList.cons (Tree.leaf 10) TreeList.nil))) + +end Foo + +/-! +Check that types and proofs are erased for both `inductive` and `structure`. +-/ + +inductive test1 : Type 1 where + | wrap : Type → 2 < 3 → test1 + deriving Repr + +structure test2 : Type 1 where + ty : Type + wrap : 2 < 3 + deriving Repr + +/-- info: test1.wrap _ _ -/ +#guard_msgs in #eval test1.wrap Nat (by simp) +/-- info: { ty := _, wrap := _ } -/ +#guard_msgs in #eval test2.mk Nat (by simp) From dcf74b0d897e31adbd7ef82f10c06769be88959c Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Wed, 8 May 2024 15:04:25 +1000 Subject: [PATCH 16/44] chore: Std -> Batteries renaming (#4108) --- .github/ISSUE_TEMPLATE/bug_report.md | 2 +- .github/workflows/pr-release.yml | 32 ++++++++++++------------- doc/BoolExpr.lean | 2 +- doc/dev/release_checklist.md | 16 ++++++------- doc/monads/states.lean | 2 +- src/Init/Data/Array/Lemmas.lean | 2 +- src/Init/Data/List/Lemmas.lean | 7 ------ src/Init/Omega/Coeffs.lean | 2 +- src/Init/Tactics.lean | 4 ++-- src/Lean/Meta/Tactic/LibrarySearch.lean | 2 +- tests/lean/librarySearch.lean | 4 ++-- 11 files changed, 34 insertions(+), 41 deletions(-) diff --git a/.github/ISSUE_TEMPLATE/bug_report.md b/.github/ISSUE_TEMPLATE/bug_report.md index 0e75ac87ec9a..2c660857033c 100644 --- a/.github/ISSUE_TEMPLATE/bug_report.md +++ b/.github/ISSUE_TEMPLATE/bug_report.md @@ -11,7 +11,7 @@ assignees: '' * [ ] Put an X between the brackets on this line if you have done all of the following: * Check that your issue is not already [filed](https://github.com/leanprover/lean4/issues). - * Reduce the issue to a minimal, self-contained, reproducible test case. Avoid dependencies to mathlib4 or std4. + * Reduce the issue to a minimal, self-contained, reproducible test case. Avoid dependencies to Mathlib or Batteries. ### Description diff --git a/.github/workflows/pr-release.yml b/.github/workflows/pr-release.yml index ef5a2f8bbaf3..1c15eda08e9a 100644 --- a/.github/workflows/pr-release.yml +++ b/.github/workflows/pr-release.yml @@ -126,11 +126,11 @@ jobs: if [ "$NIGHTLY_SHA" = "$MERGE_BASE_SHA" ]; then echo "The merge base of this PR coincides with the nightly release" - STD_REMOTE_TAGS="$(git ls-remote https://github.com/leanprover/std4.git nightly-testing-"$MOST_RECENT_NIGHTLY")" + BATTERIES_REMOTE_TAGS="$(git ls-remote https://github.com/leanprover-community/batteries.git nightly-testing-"$MOST_RECENT_NIGHTLY")" MATHLIB_REMOTE_TAGS="$(git ls-remote https://github.com/leanprover-community/mathlib4.git nightly-testing-"$MOST_RECENT_NIGHTLY")" - if [[ -n "$STD_REMOTE_TAGS" ]]; then - echo "... and Std has a 'nightly-testing-$MOST_RECENT_NIGHTLY' tag." + if [[ -n "$BATTERIES_REMOTE_TAGS" ]]; then + echo "... and Batteries has a 'nightly-testing-$MOST_RECENT_NIGHTLY' tag." MESSAGE="" if [[ -n "$MATHLIB_REMOTE_TAGS" ]]; then @@ -140,8 +140,8 @@ jobs: MESSAGE="- ❗ Mathlib CI can not be attempted yet, as the \`nightly-testing-$MOST_RECENT_NIGHTLY\` tag does not exist there yet. We will retry when you push more commits. If you rebase your branch onto \`nightly-with-mathlib\`, Mathlib CI should run now." fi else - echo "... but Std does not yet have a 'nightly-testing-$MOST_RECENT_NIGHTLY' tag." - MESSAGE="- ❗ Std CI can not be attempted yet, as the \`nightly-testing-$MOST_RECENT_NIGHTLY\` tag does not exist there yet. We will retry when you push more commits. If you rebase your branch onto \`nightly-with-mathlib\`, Std CI should run now." + echo "... but Batteries does not yet have a 'nightly-testing-$MOST_RECENT_NIGHTLY' tag." + MESSAGE="- ❗ Batteries CI can not be attempted yet, as the \`nightly-testing-$MOST_RECENT_NIGHTLY\` tag does not exist there yet. We will retry when you push more commits. If you rebase your branch onto \`nightly-with-mathlib\`, Batteries CI should run now." fi else @@ -151,7 +151,7 @@ jobs: git -C lean4.git fetch origin nightly-with-mathlib NIGHTLY_WITH_MATHLIB_SHA="$(git -C lean4.git rev-parse "origin/nightly-with-mathlib")" - MESSAGE="- ❗ Std/Mathlib CI will not be attempted unless your PR branches off the \`nightly-with-mathlib\` branch. Try \`git rebase $MERGE_BASE_SHA --onto $NIGHTLY_WITH_MATHLIB_SHA\`." + MESSAGE="- ❗ Batteries/Mathlib CI will not be attempted unless your PR branches off the \`nightly-with-mathlib\` branch. Try \`git rebase $MERGE_BASE_SHA --onto $NIGHTLY_WITH_MATHLIB_SHA\`." fi if [[ -n "$MESSAGE" ]]; then @@ -223,27 +223,27 @@ jobs: description: description, }); - # We next automatically create a Std branch using this toolchain. - # Std doesn't itself have a mechanism to report results of CI from this branch back to Lean - # Instead this is taken care of by Mathlib CI, which will fail if Std fails. + # We next automatically create a Batteries branch using this toolchain. + # Batteries doesn't itself have a mechanism to report results of CI from this branch back to Lean + # Instead this is taken care of by Mathlib CI, which will fail if Batteries fails. - name: Cleanup workspace if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' run: | sudo rm -rf ./* - # Checkout the Std repository with all branches - - name: Checkout Std repository + # Checkout the Batteries repository with all branches + - name: Checkout Batteries repository if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' uses: actions/checkout@v3 with: - repository: leanprover/std4 + repository: leanprover-community/batteries token: ${{ secrets.MATHLIB4_BOT }} ref: nightly-testing fetch-depth: 0 # This ensures we check out all tags and branches. - name: Check if tag exists if: steps.workflow-info.outputs.pullRequestNumber != '' && steps.ready.outputs.mathlib_ready == 'true' - id: check_std_tag + id: check_batteries_tag run: | git config user.name "leanprover-community-mathlib4-bot" git config user.email "leanprover-community-mathlib4-bot@users.noreply.github.com" @@ -251,7 +251,7 @@ jobs: if git ls-remote --heads --tags --exit-code origin "nightly-testing-${MOST_RECENT_NIGHTLY}" >/dev/null; then BASE="nightly-testing-${MOST_RECENT_NIGHTLY}" else - echo "This shouldn't be possible: couldn't find a 'nightly-testing-${MOST_RECENT_NIGHTLY}' tag at Std. Falling back to 'nightly-testing'." + echo "This shouldn't be possible: couldn't find a 'nightly-testing-${MOST_RECENT_NIGHTLY}' tag at Batteries. Falling back to 'nightly-testing'." BASE=nightly-testing fi @@ -268,7 +268,7 @@ jobs: else echo "Branch already exists, pushing an empty commit." git switch lean-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} - # The Std `nightly-testing` or `nightly-testing-YYYY-MM-DD` branch may have moved since this branch was created, so merge their changes. + # The Batteries `nightly-testing` or `nightly-testing-YYYY-MM-DD` branch may have moved since this branch was created, so merge their changes. # (This should no longer be possible once `nightly-testing-YYYY-MM-DD` is a tag, but it is still safe to merge.) git merge "$BASE" --strategy-option ours --no-commit --allow-unrelated-histories git commit --allow-empty -m "Trigger CI for https://github.com/leanprover/lean4/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" @@ -321,7 +321,7 @@ jobs: git switch -c lean-pr-testing-${{ steps.workflow-info.outputs.pullRequestNumber }} "$BASE" echo "leanprover/lean4-pr-releases:pr-release-${{ steps.workflow-info.outputs.pullRequestNumber }}" > lean-toolchain git add lean-toolchain - sed -i "s/require std from git \"https:\/\/github.com\/leanprover\/std4\" @ \".\+\"/require std from git \"https:\/\/github.com\/leanprover\/std4\" @ \"nightly-testing-${MOST_RECENT_NIGHTLY}\"/" lakefile.lean + sed -i "s/require batteries from git \"https:\/\/github.com\/leanprover-community\/batteries\" @ \".\+\"/require batteries from git \"https:\/\/github.com\/leanprover-community\/batteries\" @ \"nightly-testing-${MOST_RECENT_NIGHTLY}\"/" lakefile.lean git add lakefile.lean git commit -m "Update lean-toolchain for testing https://github.com/leanprover/lean4/pull/${{ steps.workflow-info.outputs.pullRequestNumber }}" else diff --git a/doc/BoolExpr.lean b/doc/BoolExpr.lean index ae9d02cbd4b9..b577ecc33e0e 100644 --- a/doc/BoolExpr.lean +++ b/doc/BoolExpr.lean @@ -1,4 +1,4 @@ -open Std +open Batteries open Lean inductive BoolExpr where diff --git a/doc/dev/release_checklist.md b/doc/dev/release_checklist.md index 90638b041acd..2e0b06b576bf 100644 --- a/doc/dev/release_checklist.md +++ b/doc/dev/release_checklist.md @@ -50,13 +50,13 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR - Create and push the tag - Merge the tag into `stable` - - [Std](https://github.com/leanprover-community/std4) + - [Batteries](https://github.com/leanprover-community/batteries) - No dependencies - Toolchain bump PR - Create and push the tag - Merge the tag into `stable` - [ProofWidgets4](https://github.com/leanprover-community/ProofWidgets4) - - Dependencies: `Std` + - Dependencies: `Batteries` - Note on versions and branches: - `ProofWidgets` uses a sequential version tagging scheme, e.g. `v0.0.29`, which does not refer to the toolchain being used. @@ -65,7 +65,7 @@ We'll use `v4.6.0` as the intended release version as a running example. - Toolchain bump PR - Create and push the tag, following the version convention of the repository - [Aesop](https://github.com/leanprover-community/aesop) - - Dependencies: `Std` + - Dependencies: `Batteries` - Toolchain bump PR including updated Lake manifest - Create and push the tag - Merge the tag into `stable` @@ -79,7 +79,7 @@ We'll use `v4.6.0` as the intended release version as a running example. - Create and push the tag - There is no `stable` branch; skip this step - [Mathlib](https://github.com/leanprover-community/mathlib4) - - Dependencies: `Aesop`, `ProofWidgets4`, `lean4checker`, `Std`, `doc-gen4`, `import-graph` + - Dependencies: `Aesop`, `ProofWidgets4`, `lean4checker`, `Batteries`, `doc-gen4`, `import-graph` - Toolchain bump PR notes: - In addition to updating the `lean-toolchain` and `lakefile.lean`, in `.github/workflows/build.yml.in` in the `lean4checker` section update the line @@ -123,8 +123,8 @@ We'll use `v4.7.0-rc1` as the intended release version in this example. - Decide which nightly release you want to turn into a release candidate. We will use `nightly-2024-02-29` in this example. -- It is essential that Std and Mathlib already have reviewed branches compatible with this nightly. - - Check that both Std and Mathlib's `bump/v4.7.0` branch contain `nightly-2024-02-29` +- It is essential that Batteries and Mathlib already have reviewed branches compatible with this nightly. + - Check that both Batteries and Mathlib's `bump/v4.7.0` branch contain `nightly-2024-02-29` in their `lean-toolchain`. - The steps required to reach that state are beyond the scope of this checklist, but see below! - Create the release branch from this nightly tag: @@ -182,7 +182,7 @@ We'll use `v4.7.0-rc1` as the intended release version in this example. - We do this for the same list of repositories as for stable releases, see above. As above, there are dependencies between these, and so the process above is iterative. It greatly helps if you can merge the `bump/v4.7.0` PRs yourself! - - For Std/Aesop/Mathlib, which maintain a `nightly-testing` branch, make sure there is a tag + - For Batteries/Aesop/Mathlib, which maintain a `nightly-testing` branch, make sure there is a tag `nightly-testing-2024-02-29` with date corresponding to the nightly used for the release (create it if not), and then on the `nightly-testing` branch `git reset --hard master`, and force push. - Make an announcement! @@ -204,7 +204,7 @@ In particular, updating the downstream repositories is significantly more work # Preparing `bump/v4.7.0` branches While not part of the release process per se, -this is a brief summary of the work that goes into updating Std/Aesop/Mathlib to new versions. +this is a brief summary of the work that goes into updating Batteries/Aesop/Mathlib to new versions. Please read https://leanprover-community.github.io/contribute/tags_and_branches.html diff --git a/doc/monads/states.lean b/doc/monads/states.lean index 0265fbe6c607..85ef83104f86 100644 --- a/doc/monads/states.lean +++ b/doc/monads/states.lean @@ -15,7 +15,7 @@ data type containing several important pieces of information. First and foremost current player, and it has a random generator. -/ -open Std (HashMap) +open Batteries (HashMap) abbrev TileIndex := Nat × Nat -- a 2D index inductive TileState where diff --git a/src/Init/Data/Array/Lemmas.lean b/src/Init/Data/Array/Lemmas.lean index 900b23b4fc6b..a3356b7410fd 100644 --- a/src/Init/Data/Array/Lemmas.lean +++ b/src/Init/Data/Array/Lemmas.lean @@ -943,7 +943,7 @@ theorem anyM_loop_iff_exists (p : α → Bool) (as : Array α) (start stop) (h : omega termination_by stop - start --- This could also be proved from `SatisfiesM_anyM_iff_exists` in `Std.Data.Array.Init.Monadic` +-- This could also be proved from `SatisfiesM_anyM_iff_exists` in `Batteries.Data.Array.Init.Monadic` theorem any_iff_exists (p : α → Bool) (as : Array α) (start stop) : any as p start stop ↔ ∃ i : Fin as.size, start ≤ i.1 ∧ i.1 < stop ∧ p as[i] := by dsimp [any, anyM, Id.run] diff --git a/src/Init/Data/List/Lemmas.lean b/src/Init/Data/List/Lemmas.lean index 00070de044fb..a694cce91e81 100644 --- a/src/Init/Data/List/Lemmas.lean +++ b/src/Init/Data/List/Lemmas.lean @@ -16,13 +16,6 @@ namespace List open Nat -/-! -# Bootstrapping theorems for lists - -These are theorems used in the definitions of `Std.Data.List.Basic` and tactics. -New theorems should be added to `Std.Data.List.Lemmas` if they are not needed by the bootstrap. --/ - attribute [simp] concat_eq_append append_assoc @[simp] theorem get?_nil : @get? α [] n = none := rfl diff --git a/src/Init/Omega/Coeffs.lean b/src/Init/Omega/Coeffs.lean index ba1ca48e2c24..57e35f6705da 100644 --- a/src/Init/Omega/Coeffs.lean +++ b/src/Init/Omega/Coeffs.lean @@ -68,7 +68,7 @@ abbrev map (f : Int → Int) (xs : Coeffs) : Coeffs := List.map f xs /-- Shim for `.enum.find?`. -/ abbrev findIdx? (f : Int → Bool) (xs : Coeffs) : Option Nat := -- List.findIdx? f xs - -- We could avoid `Std.Data.List.Basic` by using the less efficient: + -- We could avoid `Batteries.Data.List.Basic` by using the less efficient: xs.enum.find? (f ·.2) |>.map (·.1) /-- Shim for `IntList.bmod`. -/ abbrev bmod (x : Coeffs) (m : Nat) : Coeffs := IntList.bmod x m diff --git a/src/Init/Tactics.lean b/src/Init/Tactics.lean index bf58d089ca52..57dd085ed304 100644 --- a/src/Init/Tactics.lean +++ b/src/Init/Tactics.lean @@ -835,7 +835,7 @@ syntax (name := renameI) "rename_i" (ppSpace colGt binderIdent)+ : tactic /-- `repeat tac` repeatedly applies `tac` to the main goal until it fails. That is, if `tac` produces multiple subgoals, only subgoals up to the first failure will be visited. -The `Std` library provides `repeat'` which repeats separately in each subgoal. +The `Batteries` library provides `repeat'` which repeats separately in each subgoal. -/ syntax "repeat " tacticSeq : tactic macro_rules @@ -1266,7 +1266,7 @@ Optional arguments passed via a configuration argument as `solve_by_elim (config but it is often useful to change to `.reducible`, so semireducible definitions will not be unfolded when trying to apply a lemma. -See also the doc-comment for `Std.Tactic.BacktrackConfig` for the options +See also the doc-comment for `Lean.Meta.Tactic.Backtrack.BacktrackConfig` for the options `proc`, `suspend`, and `discharge` which allow further customization of `solve_by_elim`. Both `apply_assumption` and `apply_rules` are implemented via these hooks. -/ diff --git a/src/Lean/Meta/Tactic/LibrarySearch.lean b/src/Lean/Meta/Tactic/LibrarySearch.lean index f7b23466c8a2..3145c38e0786 100644 --- a/src/Lean/Meta/Tactic/LibrarySearch.lean +++ b/src/Lean/Meta/Tactic/LibrarySearch.lean @@ -160,7 +160,7 @@ An exception ID that indicates further speculation on candidate lemmas should st and current results should be returned. -/ private builtin_initialize abortSpeculationId : InternalExceptionId ← - registerInternalExceptionId `Std.Tactic.LibrarySearch.abortSpeculation + registerInternalExceptionId `Lean.Meta.LibrarySearch.abortSpeculation /-- Called to abort speculative execution in library search. diff --git a/tests/lean/librarySearch.lean b/tests/lean/librarySearch.lean index 3688efadb09d..474828fb1531 100644 --- a/tests/lean/librarySearch.lean +++ b/tests/lean/librarySearch.lean @@ -1,7 +1,7 @@ -- Enable this option for tracing: --- set_option trace.Tactic.stdLibrarySearch true +-- set_option trace.Tactic.librarySearch true -- And this option to trace all candidate lemmas before application. --- set_option trace.Tactic.stdLibrarySearch.lemmas true +-- set_option trace.Tactic.librarySearch.lemmas true -- Many of the tests here are quite volatile, -- and when changes are made to `solve_by_elim` or `exact?`, From 5814a45d449e0607d41dddcb613be70c68f5eec0 Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Wed, 8 May 2024 14:34:27 +0200 Subject: [PATCH 17/44] fix: mainModuleName should use srcSearchPath (#4066) As [reported on Zulip](https://leanprover.zulipchat.com/#narrow/stream/341532-lean4-dev/topic/Find.20references.20broken.20in.20lean.20core/near/437051935). The `mainModuleName` was being set incorrectly when browsing lean core sources, resulting in failure of cross-file server requests like "Find References". Because the `srcSearchPath` is generated asynchronously, we store it as a `Task Name` which is resolved some time before the header is finished parsing. (I don't think the `.get` here will ever block, because the srcSearchPath will be ready by the time the initial command snap is requested.) --------- Co-authored-by: Sebastian Ullrich --- src/Lean/Elab/Frontend.lean | 4 ++-- src/Lean/Language/Basic.lean | 17 ++++---------- src/Lean/Language/Lean.lean | 40 ++++++++++++++++++++++++--------- src/Lean/Server/FileWorker.lean | 33 ++++++++++++++++++--------- 4 files changed, 57 insertions(+), 37 deletions(-) diff --git a/src/Lean/Elab/Frontend.lean b/src/Lean/Elab/Frontend.lean index 0048cb4dab45..3e03bebabea8 100644 --- a/src/Lean/Elab/Frontend.lean +++ b/src/Lean/Elab/Frontend.lean @@ -154,9 +154,9 @@ def runFrontend return (s.commandState.env, !s.commandState.messages.hasErrors) - let ctx := { inputCtx with mainModuleName, opts, trustLevel } + let ctx := { inputCtx with } let processor := Language.Lean.process - let snap ← processor none ctx + let snap ← processor (fun _ => pure <| .ok { mainModuleName, opts, trustLevel }) none ctx let snaps := Language.toSnapshotTree snap snaps.runAndReport opts jsonOutput if let some ileanFileName := ileanFileName? then diff --git a/src/Lean/Language/Basic.lean b/src/Lean/Language/Basic.lean index 9ef860cb2deb..a4f4a3f952c8 100644 --- a/src/Lean/Language/Basic.lean +++ b/src/Lean/Language/Basic.lean @@ -245,17 +245,8 @@ def SnapshotTree.runAndReport (s : SnapshotTree) (opts : Options) (json := false def SnapshotTree.getAll (s : SnapshotTree) : Array Snapshot := s.forM (m := StateM _) (fun s => modify (·.push s)) |>.run #[] |>.2 -/-- Metadata that does not change during the lifetime of the language processing process. -/ -structure ModuleProcessingContext where - /-- Module name of the file being processed. -/ - mainModuleName : Name - /-- Options provided outside of the file content, e.g. on the cmdline or in the lakefile. -/ - opts : Options - /-- Kernel trust level. -/ - trustLevel : UInt32 := 0 - /-- Context of an input processing invocation. -/ -structure ProcessingContext extends ModuleProcessingContext, Parser.InputContext +structure ProcessingContext extends Parser.InputContext /-- Monad transformer holding all relevant data for processing. -/ abbrev ProcessingT m := ReaderT ProcessingContext m @@ -296,10 +287,10 @@ end Language /-- Builds a function for processing a language using incremental snapshots by passing the previous snapshot to `Language.process` on subsequent invocations. -/ -def Language.mkIncrementalProcessor (process : Option InitSnap → ProcessingM InitSnap) - (ctx : ModuleProcessingContext) : BaseIO (Parser.InputContext → BaseIO InitSnap) := do +def Language.mkIncrementalProcessor (process : Option InitSnap → ProcessingM InitSnap) : + BaseIO (Parser.InputContext → BaseIO InitSnap) := do let oldRef ← IO.mkRef none return fun ictx => do - let snap ← process (← oldRef.get) { ctx, ictx with } + let snap ← process (← oldRef.get) { ictx with } oldRef.set (some snap) return snap diff --git a/src/Lean/Language/Lean.lean b/src/Lean/Language/Lean.lean index 3f0f9fcb4178..ebcf0999bb2f 100644 --- a/src/Lean/Language/Lean.lean +++ b/src/Lean/Language/Lean.lean @@ -263,7 +263,28 @@ private def withHeaderExceptions (ex : Snapshot → α) (act : LeanProcessingT I | .error e => return ex { diagnostics := (← diagnosticsOfHeaderError e.toString) } | .ok a => return a -/-- Entry point of the Lean language processor. -/ +/-- +Result of retrieving additional metadata about the current file after parsing imports. In the +language server, these are derived from the `lake setup-file` result. On the cmdline and for similar +simple uses, these can be computed eagerly without looking at the imports. +-/ +structure SetupImportsResult where + /-- Module name of the file being processed. -/ + mainModuleName : Name + /-- Options provided outside of the file content, e.g. on the cmdline or in the lakefile. -/ + opts : Options + /-- Kernel trust level. -/ + trustLevel : UInt32 := 0 + +/-- +Entry point of the Lean language processor. + +The `setupImports` function is called after the header has been parsed and before the first command +is parsed in order to supply additional file metadata (or abort with a given terminal snapshot); see +`SetupImportsResult`. + +`old?` is a previous resulting snapshot, if any, to be reused for incremental processing. +-/ /- General notes: * For each processing function we pass in the previous state, if any, in order to reuse still-valid @@ -277,8 +298,7 @@ General notes: fast-forwarded snapshots without having to wait on tasks. -/ partial def process - (setupImports : Syntax → ProcessingT IO (Except HeaderProcessedSnapshot Options) := - fun _ => pure <| .ok {}) + (setupImports : Syntax → ProcessingT IO (Except HeaderProcessedSnapshot SetupImportsResult)) (old? : Option InitialSnapshot) : ProcessingM InitialSnapshot := do -- compute position of syntactic change once let firstDiffPos? := old?.map (·.ictx.input.firstDiffPos (← read).input) @@ -354,20 +374,18 @@ where SnapshotTask.ofIO (some ⟨0, ctx.input.endPos⟩) <| ReaderT.run (r := ctx) <| -- re-enter reader in new task withHeaderExceptions (α := HeaderProcessedSnapshot) ({ · with result? := none }) do - let opts ← match (← setupImports stx) with - | .ok opts => pure opts + let setup ← match (← setupImports stx) with + | .ok setup => pure setup | .error snap => return snap - -- override context options with file options - let opts := ctx.opts.mergeBy (fun _ _ fileOpt => fileOpt) opts -- allows `headerEnv` to be leaked, which would live until the end of the process anyway - let (headerEnv, msgLog) ← Elab.processHeader (leakEnv := true) stx opts .empty - ctx.toInputContext ctx.trustLevel + let (headerEnv, msgLog) ← Elab.processHeader (leakEnv := true) stx setup.opts .empty + ctx.toInputContext setup.trustLevel let diagnostics := (← Snapshot.Diagnostics.ofMessageLog msgLog) if msgLog.hasErrors then return { diagnostics, result? := none } - let headerEnv := headerEnv.setMainModule ctx.mainModuleName - let cmdState := Elab.Command.mkState headerEnv msgLog opts + let headerEnv := headerEnv.setMainModule setup.mainModuleName + let cmdState := Elab.Command.mkState headerEnv msgLog setup.opts let cmdState := { cmdState with infoState := { enabled := true trees := #[Elab.InfoTree.context (.commandCtx { diff --git a/src/Lean/Server/FileWorker.lean b/src/Lean/Server/FileWorker.lean index 31e237a7025f..22db77f42503 100644 --- a/src/Lean/Server/FileWorker.lean +++ b/src/Lean/Server/FileWorker.lean @@ -265,9 +265,9 @@ open Language Lean in Callback from Lean language processor after parsing imports that requests necessary information from Lake for processing imports. -/ -def setupImports (meta : DocumentMeta) (chanOut : Channel JsonRpc.Message) +def setupImports (meta : DocumentMeta) (cmdlineOpts : Options) (chanOut : Channel JsonRpc.Message) (srcSearchPathPromise : Promise SearchPath) (stx : Syntax) : - Language.ProcessingT IO (Except Language.Lean.HeaderProcessedSnapshot Options) := do + Language.ProcessingT IO (Except Language.Lean.HeaderProcessedSnapshot SetupImportsResult) := do let importsAlreadyLoaded ← importsLoadedRef.modifyGet ((·, true)) if importsAlreadyLoaded then -- As we never unload imports in the server, we should not run the code below twice in the @@ -306,27 +306,38 @@ def setupImports (meta : DocumentMeta) (chanOut : Channel JsonRpc.Message) | _ => pure () srcSearchPathPromise.resolve fileSetupResult.srcSearchPath - return .ok fileSetupResult.fileOptions + + let mainModuleName ← if let some path := System.Uri.fileUriToPath? meta.uri then + EIO.catchExceptions (h := fun _ => pure Name.anonymous) do + if let some mod ← searchModuleNameOfFileName path fileSetupResult.srcSearchPath then + pure mod + else + moduleNameOfFileName path none + else + pure Name.anonymous + + -- override cmdline options with file options + let opts := cmdlineOpts.mergeBy (fun _ _ fileOpt => fileOpt) fileSetupResult.fileOptions + + return .ok { + mainModuleName + opts + } /- Worker initialization sequence. -/ section Initialization def initializeWorker (meta : DocumentMeta) (o e : FS.Stream) (initParams : InitializeParams) (opts : Options) : IO (WorkerContext × WorkerState) := do let clientHasWidgets := initParams.initializationOptions?.bind (·.hasWidgets?) |>.getD false - let mut mainModuleName := Name.anonymous - try - if let some path := System.Uri.fileUriToPath? meta.uri then - mainModuleName ← moduleNameOfFileName path none - catch _ => pure () let maxDocVersionRef ← IO.mkRef 0 - let freshRequestIdRef ← IO.mkRef 0 + let freshRequestIdRef ← IO.mkRef (0 : Int) let chanIsProcessing ← IO.Channel.new let stickyDiagnosticsRef ← IO.mkRef ∅ let chanOut ← mkLspOutputChannel maxDocVersionRef chanIsProcessing let srcSearchPathPromise ← IO.Promise.new - let processor := Language.Lean.process (setupImports meta chanOut srcSearchPathPromise) - let processor ← Language.mkIncrementalProcessor processor { opts, mainModuleName } + let processor := Language.Lean.process (setupImports meta opts chanOut srcSearchPathPromise) + let processor ← Language.mkIncrementalProcessor processor let initSnap ← processor meta.mkInputContext let _ ← IO.mapTask (t := srcSearchPathPromise.result) fun srcSearchPath => do let importClosure := getImportClosure? initSnap From e9c302c17ec9f059d6ac6f528b8dac8712625088 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 8 May 2024 15:21:17 +0200 Subject: [PATCH 18/44] chore: bug template: point to live.lean-lang.org (#4109) https://live.lean-lang.org/#project=lean-nightly now allows users to play around with the latest lean nightly, and it seems prudent to ask them to test bug reports, if possible, there, and not just with whatever release they use. Also reformatted the descriptions to look well in a text area. Users will not see this as rendered markdown, but as plain text. --------- Co-authored-by: Sebastian Ullrich --- .github/ISSUE_TEMPLATE/bug_report.md | 16 +++++++++++----- 1 file changed, 11 insertions(+), 5 deletions(-) diff --git a/.github/ISSUE_TEMPLATE/bug_report.md b/.github/ISSUE_TEMPLATE/bug_report.md index 2c660857033c..8bb4dc9c19a3 100644 --- a/.github/ISSUE_TEMPLATE/bug_report.md +++ b/.github/ISSUE_TEMPLATE/bug_report.md @@ -9,9 +9,15 @@ assignees: '' ### Prerequisites -* [ ] Put an X between the brackets on this line if you have done all of the following: - * Check that your issue is not already [filed](https://github.com/leanprover/lean4/issues). - * Reduce the issue to a minimal, self-contained, reproducible test case. Avoid dependencies to Mathlib or Batteries. +Please put an X between the brackets as you perform the following steps: + +* [ ] Check that your issue is not already filed: + https://github.com/leanprover/lean4/issues +* [ ] Reduce the issue to a minimal, self-contained, reproducible test case. + Avoid dependencies to Mathlib or Batteries. +* [ ] Test your test case against the latest nightly release, for example on + https://live.lean-lang.org/#project=lean-nightly + (You can also use the settings there to switch to “Lean nightly”) ### Description @@ -33,8 +39,8 @@ assignees: '' ### Versions -[Output of `#eval Lean.versionString` or of `lean --version` in the folder that the issue occured in] -[OS version] +[Output of `#eval Lean.versionString`] +[OS version, if not using live.lean-lang.org.] ### Additional Information From 227e8617190f1685445616e0848aa556faa35b1c Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Wed, 8 May 2024 17:12:47 +0200 Subject: [PATCH 19/44] refactor: make 1 % n reduce without well-founded recursion (#4098) this is in preparation for #4061. Once that lands, `1 % 42 = 1` will no longer hold definitionally (at least not without an ungly `unseal Nat.modCore in` around). This affects mathlib in a few places, essentially every time a `1 : Fin (n+1)` literal is written. So this extends the existing special case for `0 % n = 0` to `1 % n`. --- src/Init/Data/Int/DivModLemmas.lean | 6 ++++-- src/Init/Data/Nat/Div.lean | 24 +++++++++++++++--------- 2 files changed, 19 insertions(+), 11 deletions(-) diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index 2679a9ace5a2..d532d40a4272 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -178,7 +178,7 @@ theorem fdiv_eq_div {a b : Int} (Ha : 0 ≤ a) (Hb : 0 ≤ b) : fdiv a b = div a @[simp] theorem mod_zero : ∀ a : Int, mod a 0 = a | ofNat _ => congrArg ofNat <| Nat.mod_zero _ - | -[_+1] => rfl + | -[_+1] => congrArg (fun n => -ofNat n) <| Nat.mod_zero _ @[simp] theorem zero_fmod (b : Int) : fmod 0 b = 0 := by cases b <;> rfl @@ -225,7 +225,9 @@ theorem mod_add_div : ∀ a b : Int, mod a b + b * (a.div b) = a | 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], 0 => by + show -(↑((succ m) % 0) : Int) + 0 * -↑(succ m / 0) = -↑(succ m) + rw [Nat.mod_zero, Int.zero_mul, Int.add_zero] | -[m+1], ofNat n => by show -(↑((succ m) % n) : Int) + ↑n * -↑(succ m / n) = -↑(succ m) rw [Int.mul_neg, ← Int.neg_add] diff --git a/src/Init/Data/Nat/Div.lean b/src/Init/Data/Nat/Div.lean index 8bd98b8d5215..ae531772cb0a 100644 --- a/src/Init/Data/Nat/Div.lean +++ b/src/Init/Data/Nat/Div.lean @@ -82,22 +82,28 @@ decreasing_by apply div_rec_lemma; assumption @[extern "lean_nat_mod"] protected def mod : @& Nat → @& Nat → Nat - /- This case is not needed mathematically as the case below is equal to it; however, it makes - `0 % n = 0` true definitionally rather than just propositionally. - This property is desirable for `Fin n`, as it means `(ofNat 0 : Fin n).val = 0` by definition. - Primarily, this is valuable because mathlib in Lean3 assumed this was true definitionally, and so - keeping this definitional equality makes mathlib easier to port to mathlib4. -/ + /- These four cases are not needed mathematically, they are just special cases of the + general case. However, it makes `0 % n = 0` etc. true definitionally rather than just + propositionally. + This property is desirable for `Fin n` literals, as it means `(ofNat 0 : Fin n).val = 0` by + definition. This was true in lean3 and it simplified things for mathlib if it remains true. -/ | 0, _ => 0 - | x@(_ + 1), y => Nat.modCore x y + | 1, 0 => 0 + | 1, 1 => 0 + | 1, (_+2) => 1 + | x@(_ + 2), y => Nat.modCore x y instance instMod : Mod Nat := ⟨Nat.mod⟩ protected theorem modCore_eq_mod (x y : Nat) : Nat.modCore x y = x % y := by - cases x with - | zero => + match x, y with + | 0, y => rw [Nat.modCore] exact if_neg fun ⟨hlt, hle⟩ => Nat.lt_irrefl _ (Nat.lt_of_lt_of_le hlt hle) - | succ x => rfl + | 1, 0 => rw [Nat.modCore]; rfl + | 1, 1 => rw [Nat.modCore, Nat.modCore]; rfl + | 1, (_+2) => rw [Nat.modCore]; rfl + | (_ + 2), _ => rfl 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] From d7c69205508018bd66edaa5d9f4a2fa3e4f86ec3 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Wed, 8 May 2024 17:38:04 +0200 Subject: [PATCH 20/44] chore: remove dead function`preprocessLevels` (#4112) --- src/Lean/Meta/SynthInstance.lean | 12 ------------ 1 file changed, 12 deletions(-) diff --git a/src/Lean/Meta/SynthInstance.lean b/src/Lean/Meta/SynthInstance.lean index ca6cfd5b26c0..6a82fe7c5767 100644 --- a/src/Lean/Meta/SynthInstance.lean +++ b/src/Lean/Meta/SynthInstance.lean @@ -664,18 +664,6 @@ private def preprocess (type : Expr) : MetaM Expr := let type ← whnf type mkForallFVars xs type -private def preprocessLevels (us : List Level) : MetaM (List Level × Bool) := do - let mut r := #[] - let mut modified := false - for u in us do - let u ← instantiateLevelMVars u - if u.hasMVar then - r := r.push (← mkFreshLevelMVar) - modified := true - else - r := r.push u - return (r.toList, modified) - private partial def preprocessArgs (type : Expr) (i : Nat) (args : Array Expr) (outParamsPos : Array Nat) : MetaM (Array Expr) := do if h : i < args.size then let type ← whnf type From ec872834655e55835e010347e9adacc85d7307f9 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 9 May 2024 02:58:43 +0200 Subject: [PATCH 21/44] perf: use `withSynthesize` when elaborating `let`/`have` type (#4096) closes #4051 cc @semorrison --- src/Lean/Elab/Binders.lean | 25 ++++++- src/Lean/Elab/BuiltinNotation.lean | 6 +- src/Lean/Elab/Extra.lean | 3 +- src/Lean/Elab/Inductive.lean | 2 +- src/Lean/Elab/StructInst.lean | 2 +- src/Lean/Elab/SyntheticMVars.lean | 90 ++++++++++++++++++------- src/Lean/Elab/Tactic/BuiltinTactic.lean | 2 +- src/Lean/Elab/Tactic/ElabTerm.lean | 2 +- src/Lean/Elab/Tactic/Induction.lean | 2 +- src/Lean/Elab/Tactic/Meta.lean | 2 +- src/Lean/Elab/Tactic/Simp.lean | 4 +- tests/lean/run/4051.lean | 21 ++++++ tests/lean/run/may_postpone_tc.lean | 11 +++ 13 files changed, 132 insertions(+), 40 deletions(-) create mode 100644 tests/lean/run/4051.lean create mode 100644 tests/lean/run/may_postpone_tc.lean diff --git a/src/Lean/Elab/Binders.lean b/src/Lean/Elab/Binders.lean index 0c2277b939a4..9976d9141e26 100644 --- a/src/Lean/Elab/Binders.lean +++ b/src/Lean/Elab/Binders.lean @@ -7,6 +7,7 @@ prelude import Lean.Elab.Quotation.Precheck import Lean.Elab.Term import Lean.Elab.BindersUtil +import Lean.Elab.SyntheticMVars import Lean.Elab.PreDefinition.WF.TerminationHint namespace Lean.Elab.Term @@ -646,7 +647,29 @@ def elabLetDeclAux (id : Syntax) (binders : Array Syntax) (typeStx : Syntax) (va (expectedType? : Option Expr) (useLetExpr : Bool) (elabBodyFirst : Bool) (usedLetOnly : Bool) : TermElabM Expr := do let (type, val, binders) ← elabBindersEx binders fun xs => do let (binders, fvars) := xs.unzip - let type ← elabType typeStx + /- + We use `withSynthesize` to ensure that any postponed elaboration problem + and nested tactics in `type` are resolved before elaborating `val`. + Resolved: we want to avoid synthethic opaque metavariables in `type`. + Recall that this kind of metavariable is non-assignable, and `isDefEq` + may waste a lot of time unfolding declarations before failing. + See issue #4051 for an example. + + Here is the analysis for issue #4051. + - Given `have x : type := value; body`, we were previously elaborating `value` even + if `type` contained postponed elaboration problems. + - Moreover, the metavariables in `type` corresponding to postponed elaboration + problems cannot be assigned by `isDefEq` since the elaborator is supposed to assign them. + - Then, when checking whether type of `value` is definitionally equal to `type`, + a very long-time was spent unfolding a bunch of declarations before it failed. + In #4051, it was unfolding `Array.swaps` which is defined by well-founded recursion. + After the failure, the elaborator inserted a postponed coercion + that would be resolved later as soon as the types don't have unassigned metavariables. + + We use `postpone := .partial` to allow type class (TC) resolution problems to be postponed + Recall that TC resolution does **not** produce synthetic opaque metavariables. + -/ + let type ← withSynthesize (postpone := .partial) <| elabType typeStx registerCustomErrorIfMVar type typeStx "failed to infer 'let' declaration type" if elabBodyFirst then let type ← mkForallFVars fvars type diff --git a/src/Lean/Elab/BuiltinNotation.lean b/src/Lean/Elab/BuiltinNotation.lean index 56cdc5e0dfaa..cca120f25ca0 100644 --- a/src/Lean/Elab/BuiltinNotation.lean +++ b/src/Lean/Elab/BuiltinNotation.lean @@ -98,7 +98,7 @@ open Meta show Nat from 0 ``` -/ - let type ← withSynthesize (mayPostpone := true) do + let type ← withSynthesize (postpone := .yes) do let type ← elabType type if let some expectedType := expectedType? then -- Recall that a similar approach is used when elaborating applications @@ -314,11 +314,11 @@ where @[builtin_term_elab typeAscription] def elabTypeAscription : TermElab | `(($e : $type)), _ => do - let type ← withSynthesize (mayPostpone := true) <| elabType type + let type ← withSynthesize (postpone := .yes) <| elabType type let e ← elabTerm e type ensureHasType type e | `(($e :)), expectedType? => do - let e ← withSynthesize (mayPostpone := false) <| elabTerm e none + let e ← withSynthesize (postpone := .no) <| elabTerm e none ensureHasType expectedType? e | _, _ => throwUnsupportedSyntax diff --git a/src/Lean/Elab/Extra.lean b/src/Lean/Elab/Extra.lean index c4acb68e1238..abd082cb5194 100644 --- a/src/Lean/Elab/Extra.lean +++ b/src/Lean/Elab/Extra.lean @@ -188,7 +188,7 @@ private partial def toTree (s : Syntax) : TermElabM Tree := do the macro declaration names in the `op` nodes. -/ let result ← go s - synthesizeSyntheticMVars (mayPostpone := true) + synthesizeSyntheticMVars (postpone := .yes) return result where go (s : Syntax) := do @@ -486,7 +486,6 @@ def elabBinRelCore (noProp : Bool) (stx : Syntax) (expectedType? : Option Expr) | some f => withSynthesizeLight do /- We used to use `withSynthesize (mayPostpone := true)` here instead of `withSynthesizeLight` here. - Recall that `withSynthesizeLight` is equivalent to `withSynthesize (mayPostpone := true) (synthesizeDefault := false)`. It seems too much to apply default instances at binary relations. For example, we cannot elaborate ``` def as : List Int := [-1, 2, 0, -3, 4] diff --git a/src/Lean/Elab/Inductive.lean b/src/Lean/Elab/Inductive.lean index cc7e2d688d88..f73200252410 100644 --- a/src/Lean/Elab/Inductive.lean +++ b/src/Lean/Elab/Inductive.lean @@ -324,7 +324,7 @@ private def elabCtors (indFVars : Array Expr) (indFVar : Expr) (params : Array E | some ctorType => let type ← Term.elabType ctorType trace[Elab.inductive] "elabType {ctorView.declName} : {type} " - Term.synthesizeSyntheticMVars (mayPostpone := true) + Term.synthesizeSyntheticMVars (postpone := .yes) let type ← instantiateMVars type let type ← checkParamOccs type forallTelescopeReducing type fun _ resultingType => do diff --git a/src/Lean/Elab/StructInst.lean b/src/Lean/Elab/StructInst.lean index 00ee772bfd1e..65290fc41eef 100644 --- a/src/Lean/Elab/StructInst.lean +++ b/src/Lean/Elab/StructInst.lean @@ -939,7 +939,7 @@ private def elabStructInstAux (stx : Syntax) (expectedType? : Option Expr) (sour TODO: investigate whether this design decision may have unintended side effects or produce confusing behavior. -/ - let { val := r, struct, instMVars } ← withSynthesize (mayPostpone := true) <| elabStruct struct expectedType? + let { val := r, struct, instMVars } ← withSynthesize (postpone := .yes) <| elabStruct struct expectedType? trace[Elab.struct] "before propagate {r}" DefaultFields.propagate struct synthesizeAppInstMVars instMVars r diff --git a/src/Lean/Elab/SyntheticMVars.lean b/src/Lean/Elab/SyntheticMVars.lean index a39d02766165..f605436c94c6 100644 --- a/src/Lean/Elab/SyntheticMVars.lean +++ b/src/Lean/Elab/SyntheticMVars.lean @@ -288,6 +288,32 @@ private def processPostponedUniverseContraints : TermElabM Unit := do private def markAsResolved (mvarId : MVarId) : TermElabM Unit := modify fun s => { s with syntheticMVars := s.syntheticMVars.erase mvarId } +/-- +Auxiliary type for `synthesizeSyntheticMVars`. It specifies +whether pending synthetic metavariables can be postponed or not. +-/ +inductive PostponeBehavior where + /-- + Any kind of pending synthetic metavariable can be postponed. + Universe constrains may also be postponed. + -/ + | yes + /-- + Pending synthetic metavariables cannot be postponed. + -/ + | no + /-- + Synthectic metavariables associated with type class resolution can be postponed. + Motivation: this kind of metavariable are not synthethic opaque, and can be assigned by `isDefEq`. + Unviverse constraints can also be postponed. + -/ + | «partial» + deriving Inhabited, Repr, BEq + +def PostponeBehavior.ofBool : Bool → PostponeBehavior + | true => .yes + | false => .no + mutual /-- @@ -321,7 +347,7 @@ mutual -- view even though it is synthetic while a node like `tacticCode` never is (#1990) withTacticInfoContext tacticCode[0] do evalTactic code - synthesizeSyntheticMVars (mayPostpone := false) + synthesizeSyntheticMVars (postpone := .no) unless remainingGoals.isEmpty do if report then reportUnsolvedGoals remainingGoals @@ -388,25 +414,27 @@ mutual return numSyntheticMVars != remainingPendingMVars.length /-- - Try to process pending synthetic metavariables. If `mayPostpone == false`, - then `pendingMVars` is `[]` after executing this method. + Try to process pending synthetic metavariables. + + If `postpone == .no`,then `pendingMVars` is `[]` after executing this method. + If `postpone == .partial`, then `pendingMVars` contains only `.tc` and `.coe` kinds. It keeps executing `synthesizeSyntheticMVarsStep` while progress is being made. - If `mayPostpone == false`, then it applies default instances to `SyntheticMVarKind.typeClass` (if available) + If `postpone != .yes`, then it applies default instances to `SyntheticMVarKind.typeClass` (if available) metavariables that are still unresolved, and then tries to resolve metavariables - with `mayPostpone == false`. That is, we force them to produce error messages and/or commit to - a "best option". If, after that, we still haven't made progress, we report "stuck" errors. + with `postponeOnError == false`. That is, we force them to produce error messages and/or commit to + a "best option". If, after that, we still haven't made progress, we report "stuck" errors If `postpone == .no`. Remark: we set `ignoreStuckTC := true` when elaborating `simp` arguments. Then, pending TC problems become implicit parameters for the simp theorem. -/ - partial def synthesizeSyntheticMVars (mayPostpone := true) (ignoreStuckTC := false) : TermElabM Unit := do + partial def synthesizeSyntheticMVars (postpone := PostponeBehavior.yes) (ignoreStuckTC := false) : TermElabM Unit := do let rec loop (_ : Unit) : TermElabM Unit := do withRef (← getSomeSyntheticMVarsRef) <| withIncRecDepth do unless (← get).pendingMVars.isEmpty do if ← synthesizeSyntheticMVarsStep (postponeOnError := false) (runTactics := false) then loop () - else if !mayPostpone then + else if postpone != .yes then /- Resume pending metavariables with "elaboration postponement" disabled. We postpone elaboration errors in this step by setting `postponeOnError := true`. Example: @@ -431,48 +459,58 @@ mutual loop () else if ← synthesizeSyntheticMVarsStep (postponeOnError := false) (runTactics := true) then loop () - else + else if postpone == .no then reportStuckSyntheticMVars ignoreStuckTC loop () - unless mayPostpone do + if postpone == .no then processPostponedUniverseContraints end def synthesizeSyntheticMVarsNoPostponing (ignoreStuckTC := false) : TermElabM Unit := - synthesizeSyntheticMVars (mayPostpone := false) (ignoreStuckTC := ignoreStuckTC) + synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := ignoreStuckTC) /-- Keep invoking `synthesizeUsingDefault` until it returns false. -/ private partial def synthesizeUsingDefaultLoop : TermElabM Unit := do if (← synthesizeUsingDefault) then - synthesizeSyntheticMVars (mayPostpone := true) + synthesizeSyntheticMVars (postpone := .yes) synthesizeUsingDefaultLoop def synthesizeSyntheticMVarsUsingDefault : TermElabM Unit := do - synthesizeSyntheticMVars (mayPostpone := true) + synthesizeSyntheticMVars (postpone := .yes) synthesizeUsingDefaultLoop -private partial def withSynthesizeImp {α} (k : TermElabM α) (mayPostpone : Bool) (synthesizeDefault : Bool) : TermElabM α := do +private partial def withSynthesizeImp (k : TermElabM α) (postpone : PostponeBehavior) : TermElabM α := do + let pendingMVarsSaved := (← get).pendingMVars + modify fun s => { s with pendingMVars := [] } + try + let a ← k + synthesizeSyntheticMVars (postpone := postpone) + if postpone == .yes then + synthesizeUsingDefaultLoop + return a + finally + modify fun s => { s with pendingMVars := s.pendingMVars ++ pendingMVarsSaved } + +/-- + Execute `k`, and synthesize pending synthetic metavariables created while executing `k` are solved. + If `mayPostpone == false`, then all of them must be synthesized. + Remark: even if `mayPostpone == true`, the method still uses `synthesizeUsingDefault` -/ +@[inline] def withSynthesize [MonadFunctorT TermElabM m] [Monad m] (k : m α) (postpone := PostponeBehavior.no) : m α := + monadMap (m := TermElabM) (withSynthesizeImp · postpone) k + +private partial def withSynthesizeLightImp (k : TermElabM α) : TermElabM α := do let pendingMVarsSaved := (← get).pendingMVars modify fun s => { s with pendingMVars := [] } try let a ← k - synthesizeSyntheticMVars mayPostpone - if mayPostpone && synthesizeDefault then - synthesizeUsingDefaultLoop + synthesizeSyntheticMVars (postpone := .yes) return a finally modify fun s => { s with pendingMVars := s.pendingMVars ++ pendingMVarsSaved } -/-- - Execute `k`, and synthesize pending synthetic metavariables created while executing `k` are solved. - If `mayPostpone == false`, then all of them must be synthesized. - Remark: even if `mayPostpone == true`, the method still uses `synthesizeUsingDefault` -/ -@[inline] def withSynthesize [MonadFunctorT TermElabM m] [Monad m] (k : m α) (mayPostpone := false) : m α := - monadMap (m := TermElabM) (withSynthesizeImp · mayPostpone (synthesizeDefault := true)) k - -/-- Similar to `withSynthesize`, but sets `mayPostpone` to `true`, and do not use `synthesizeUsingDefault` -/ +/-- Similar to `withSynthesize`, but uses `postpone := .true`, does not use use `synthesizeUsingDefault` -/ @[inline] def withSynthesizeLight [MonadFunctorT TermElabM m] [Monad m] (k : m α) : m α := - monadMap (m := TermElabM) (withSynthesizeImp · (mayPostpone := true) (synthesizeDefault := false)) k + monadMap (m := TermElabM) (withSynthesizeLightImp ·) k /-- Elaborate `stx`, and make sure all pending synthetic metavariables created while elaborating `stx` are solved. -/ def elabTermAndSynthesize (stx : Syntax) (expectedType? : Option Expr) : TermElabM Expr := diff --git a/src/Lean/Elab/Tactic/BuiltinTactic.lean b/src/Lean/Elab/Tactic/BuiltinTactic.lean index ea40cb6f0fd3..f7e3db804f90 100644 --- a/src/Lean/Elab/Tactic/BuiltinTactic.lean +++ b/src/Lean/Elab/Tactic/BuiltinTactic.lean @@ -270,7 +270,7 @@ where pure (fvarId, [mvarId]) if let some typeStx := typeStx? then withMainContext do - let type ← Term.withSynthesize (mayPostpone := true) <| Term.elabType typeStx + let type ← Term.withSynthesize (postpone := .yes) <| Term.elabType typeStx let fvar := mkFVar fvarId let fvarType ← inferType fvar unless (← isDefEqGuarded type fvarType) do diff --git a/src/Lean/Elab/Tactic/ElabTerm.lean b/src/Lean/Elab/Tactic/ElabTerm.lean index a1171cd8f30e..1d686fe6a567 100644 --- a/src/Lean/Elab/Tactic/ElabTerm.lean +++ b/src/Lean/Elab/Tactic/ElabTerm.lean @@ -29,7 +29,7 @@ def runTermElab (k : TermElabM α) (mayPostpone := false) : TacticM α := do else Term.withoutErrToSorry go where - go := k <* Term.synthesizeSyntheticMVars (mayPostpone := mayPostpone) + go := k <* Term.synthesizeSyntheticMVars (postpone := .ofBool mayPostpone) /-- Elaborate `stx` in the current `MVarContext`. If given, the `expectedType` will be used to help elaboration but not enforced (use `elabTermEnsuringType` to enforce an expected type). -/ diff --git a/src/Lean/Elab/Tactic/Induction.lean b/src/Lean/Elab/Tactic/Induction.lean index e74577c6aea1..99caf6ec1594 100644 --- a/src/Lean/Elab/Tactic/Induction.lean +++ b/src/Lean/Elab/Tactic/Induction.lean @@ -524,7 +524,7 @@ private def elabTermForElim (stx : Syntax) : TermElabM Expr := do return e Term.withoutErrToSorry <| Term.withoutHeedElabAsElim do let e ← Term.elabTerm stx none (implicitLambda := false) - Term.synthesizeSyntheticMVars (mayPostpone := false) (ignoreStuckTC := true) + Term.synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true) let e ← instantiateMVars e let e := e.eta if e.hasMVar then diff --git a/src/Lean/Elab/Tactic/Meta.lean b/src/Lean/Elab/Tactic/Meta.lean index bb7e5a8cf5df..c6be92112e92 100644 --- a/src/Lean/Elab/Tactic/Meta.lean +++ b/src/Lean/Elab/Tactic/Meta.lean @@ -14,7 +14,7 @@ open Term def runTactic (mvarId : MVarId) (tacticCode : Syntax) (ctx : Context := {}) (s : State := {}) : MetaM (List MVarId × State) := do instantiateMVarDeclMVars mvarId let go : TermElabM (List MVarId) := - withSynthesize (mayPostpone := false) do Tactic.run mvarId (Tactic.evalTactic tacticCode *> Tactic.pruneSolvedGoals) + withSynthesize do Tactic.run mvarId (Tactic.evalTactic tacticCode *> Tactic.pruneSolvedGoals) go.run ctx s end Lean.Elab diff --git a/src/Lean/Elab/Tactic/Simp.lean b/src/Lean/Elab/Tactic/Simp.lean index 6c3a89d4ed27..a87a2ac23bc0 100644 --- a/src/Lean/Elab/Tactic/Simp.lean +++ b/src/Lean/Elab/Tactic/Simp.lean @@ -46,7 +46,7 @@ def tacticToDischarge (tacticCode : Syntax) : TacticM (IO.Ref Term.State × Simp So, we must not save references to them at `Term.State`. -/ withoutModifyingStateWithInfoAndMessages do - Term.withSynthesize (mayPostpone := false) do + Term.withSynthesize (postpone := .no) do Term.runTactic (report := false) mvar.mvarId! tacticCode let result ← instantiateMVars mvar if result.hasExprMVar then @@ -121,7 +121,7 @@ private def addDeclToUnfoldOrTheorem (thms : SimpTheorems) (id : Origin) (e : Ex private def addSimpTheorem (thms : SimpTheorems) (id : Origin) (stx : Syntax) (post : Bool) (inv : Bool) : TermElabM SimpTheorems := do let (levelParams, proof) ← Term.withoutModifyingElabMetaStateWithInfo <| withRef stx <| Term.withoutErrToSorry do let e ← Term.elabTerm stx none - Term.synthesizeSyntheticMVars (mayPostpone := false) (ignoreStuckTC := true) + Term.synthesizeSyntheticMVars (postpone := .no) (ignoreStuckTC := true) let e ← instantiateMVars e let e := e.eta if e.hasMVar then diff --git a/tests/lean/run/4051.lean b/tests/lean/run/4051.lean new file mode 100644 index 000000000000..1f877b2ad6ba --- /dev/null +++ b/tests/lean/run/4051.lean @@ -0,0 +1,21 @@ +def Array.swaps (a : Array α) : List (Fin a.size × Fin a.size) → Array α + | [] => a + | (i, j) :: ijs => + have : (a.swap i j).size = a.size := a.size_swap _ _ + + swaps (a.swap i j) (ijs.map (fun p => ⟨⟨p.1.1, this.symm ▸ p.1.2⟩, ⟨p.2.1, this.symm ▸ p.2.2⟩⟩)) +termination_by l => l.length + +set_option maxHeartbeats 1000 in +theorem Array.swaps_cancel (a : Array α) (l : List (Fin a.size × Fin a.size)) : a.swaps (l ++ l.reverse) = a := + match l with + | [] => sorry + | c :: cs => + + have h : a.size = (a.swaps [c]).size := sorry + + have ih1 : ((a.swaps [c]).swaps ((h ▸ cs) ++ (h ▸ cs).reverse)) = (a.swaps [c]) := + swaps_cancel (a.swaps [c]) (h ▸ cs) + sorry +termination_by l.length +decreasing_by sorry diff --git a/tests/lean/run/may_postpone_tc.lean b/tests/lean/run/may_postpone_tc.lean new file mode 100644 index 000000000000..a36a7e498ccb --- /dev/null +++ b/tests/lean/run/may_postpone_tc.lean @@ -0,0 +1,11 @@ +import Lean + +open Lean +def f (xs : List String) : CoreM (Array String) := do + let mut found : RBMap _ _ compare := {} + let mut result := #[] + for x in xs do + unless found.contains x do + result := result.push x + found := found.insert x () + return result From fe7b96d8a0cf85e1ea81791ac9d673f84c9e866d Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Thu, 9 May 2024 14:52:09 +1000 Subject: [PATCH 22/44] fix: generate deprecation warnings for dot notation (#3969) Fixes #3270 by moving the deprecation check from `Lean.Elab.Term.mkConsts` to `Lean.Elab.Term.mkConst`, so `Lean.Elab.Term.mkBaseProjections`, `.elabAppLValsAux`, `.elabAppFn`, and `.elabForIn` also hit the check. Not all of these really need to hit the check, so I'll run `!bench` to see if it's a problem. --- src/Init/Data/Array/Subarray.lean | 2 +- src/Init/Data/BitVec/Lemmas.lean | 2 +- src/Init/Data/Stream.lean | 2 +- src/Lean/Data/RBTree.lean | 2 +- src/Lean/Elab/Term.lean | 2 +- tests/lean/deprecated.lean | 6 ++++++ 6 files changed, 11 insertions(+), 5 deletions(-) diff --git a/src/Init/Data/Array/Subarray.lean b/src/Init/Data/Array/Subarray.lean index 682fbb204bab..f38496ec4c5e 100644 --- a/src/Init/Data/Array/Subarray.lean +++ b/src/Init/Data/Array/Subarray.lean @@ -22,7 +22,7 @@ abbrev Subarray.as (s : Subarray α) : Array α := s.array theorem Subarray.h₁ (s : Subarray α) : s.start ≤ s.stop := s.start_le_stop @[deprecated Subarray.stop_le_array_size] -theorem Subarray.h₂ (s : Subarray α) : s.stop ≤ s.as.size := s.stop_le_array_size +theorem Subarray.h₂ (s : Subarray α) : s.stop ≤ s.array.size := s.stop_le_array_size namespace Subarray diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 27f3632e23d1..5a4fbcf1d9c7 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -897,7 +897,7 @@ theorem sub_toAdd {n} (x y : BitVec n) : x - y = x + - y := by theorem add_sub_cancel (x y : BitVec w) : x + y - y = x := by apply eq_of_toNat_eq - have y_toNat_le := Nat.le_of_lt y.toNat_lt + have y_toNat_le := Nat.le_of_lt y.isLt rw [toNat_sub, toNat_add, Nat.mod_add_mod, Nat.add_assoc, ← Nat.add_sub_assoc y_toNat_le, Nat.add_sub_cancel_left, Nat.add_mod_right, toNat_mod_cancel] diff --git a/src/Init/Data/Stream.lean b/src/Init/Data/Stream.lean index f045327471fc..54cc28b13ebf 100644 --- a/src/Init/Data/Stream.lean +++ b/src/Init/Data/Stream.lean @@ -94,7 +94,7 @@ instance : Stream (Subarray α) α where next? s := if h : s.start < s.stop then have : s.start + 1 ≤ s.stop := Nat.succ_le_of_lt h - some (s.as.get ⟨s.start, Nat.lt_of_lt_of_le h s.stop_le_array_size⟩, + some (s.array.get ⟨s.start, Nat.lt_of_lt_of_le h s.stop_le_array_size⟩, { s with start := s.start + 1, start_le_stop := this }) else none diff --git a/src/Lean/Data/RBTree.lean b/src/Lean/Data/RBTree.lean index a2fde26f33b8..f749d80f8ff4 100644 --- a/src/Lean/Data/RBTree.lean +++ b/src/Lean/Data/RBTree.lean @@ -100,7 +100,7 @@ def fromArray (l : Array α) (cmp : α → α → Ordering) : RBTree α cmp := RBMap.any t (fun a _ => p a) def subset (t₁ t₂ : RBTree α cmp) : Bool := - t₁.all fun a => (t₂.find? a).toBool + t₁.all fun a => (t₂.find? a).isSome def seteq (t₁ t₂ : RBTree α cmp) : Bool := subset t₁ t₂ && subset t₂ t₁ diff --git a/src/Lean/Elab/Term.lean b/src/Lean/Elab/Term.lean index ce929262f801..ee1014f94442 100644 --- a/src/Lean/Elab/Term.lean +++ b/src/Lean/Elab/Term.lean @@ -1634,6 +1634,7 @@ def isLetRecAuxMVar (mvarId : MVarId) : TermElabM Bool := do Remark: fresh universe metavariables are created if the constant has more universe parameters than `explicitLevels`. -/ def mkConst (constName : Name) (explicitLevels : List Level := []) : TermElabM Expr := do + Linter.checkDeprecated constName -- TODO: check is occurring too early if there are multiple alternatives. Fix if it is not ok in practice let cinfo ← getConstInfo constName if explicitLevels.length > cinfo.levelParams.length then throwError "too many explicit universe levels for '{constName}'" @@ -1645,7 +1646,6 @@ def mkConst (constName : Name) (explicitLevels : List Level := []) : TermElabM E private def mkConsts (candidates : List (Name × List String)) (explicitLevels : List Level) : TermElabM (List (Expr × List String)) := do candidates.foldlM (init := []) fun result (declName, projs) => do -- TODO: better support for `mkConst` failure. We may want to cache the failures, and report them if all candidates fail. - Linter.checkDeprecated declName -- TODO: check is occurring too early if there are multiple alternatives. Fix if it is not ok in practice let const ← mkConst declName explicitLevels return (const, projs) :: result diff --git a/tests/lean/deprecated.lean b/tests/lean/deprecated.lean index d92bde5d0579..4c33c854b8eb 100644 --- a/tests/lean/deprecated.lean +++ b/tests/lean/deprecated.lean @@ -31,3 +31,9 @@ def f4 (x : Nat) := x + 1 set_option linter.deprecated false in #eval f2 0 + 1 #eval f4 0 + 1 + +@[deprecated] def Nat.z (x : Nat) := x + 1 + +/-- warning: `Nat.z` has been deprecated -/ +#guard_msgs in +example (n : Nat) : n.z = n + 1 := rfl From 6a040ab068dd7614da3e71239709ddebabffd7bf Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Thu, 9 May 2024 19:44:19 +0200 Subject: [PATCH 23/44] feat: propagate `maxHeartbeats` to kernel (#4113) Co-authored-by: Sebastian Ullrich --- src/Lean.lean | 1 + src/Lean/AddDecl.lean | 31 ++++++++++++++++++++++++ src/Lean/Compiler/InitAttr.lean | 1 + src/Lean/Compiler/Old.lean | 5 ---- src/Lean/CoreM.lean | 13 ---------- src/Lean/Elab/BuiltinCommand.lean | 2 +- src/Lean/Environment.lean | 2 +- src/Lean/Meta/Closure.lean | 1 + src/Lean/Meta/Constructions.lean | 1 + src/Lean/Meta/Eqns.lean | 1 + src/Lean/Meta/Eval.lean | 1 + src/Lean/Meta/SizeOf.lean | 1 + src/Lean/Meta/Tactic/AuxLemma.lean | 1 + src/Lean/Replay.lean | 3 ++- src/kernel/environment.cpp | 3 ++- src/runtime/interrupt.cpp | 2 +- src/runtime/interrupt.h | 4 +-- tests/lean/partialIssue.lean | 2 +- tests/lean/run/kernel_maxheartbeats.lean | 22 +++++++++++++++++ tests/lean/run/structInstFast.lean | 11 ++++++--- 20 files changed, 79 insertions(+), 29 deletions(-) create mode 100644 src/Lean/AddDecl.lean create mode 100644 tests/lean/run/kernel_maxheartbeats.lean diff --git a/src/Lean.lean b/src/Lean.lean index e026559b8ba7..5da1bb9c2435 100644 --- a/src/Lean.lean +++ b/src/Lean.lean @@ -37,3 +37,4 @@ import Lean.Log import Lean.Linter import Lean.SubExpr import Lean.LabelAttribute +import Lean.AddDecl diff --git a/src/Lean/AddDecl.lean b/src/Lean/AddDecl.lean new file mode 100644 index 000000000000..4a203f01af7f --- /dev/null +++ b/src/Lean/AddDecl.lean @@ -0,0 +1,31 @@ +/- +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.CoreM + +namespace Lean + +def Environment.addDecl (env : Environment) (opts : Options) (decl : Declaration) : Except KernelException Environment := + addDeclCore env (Core.getMaxHeartbeats opts).toUSize decl + +def Environment.addAndCompile (env : Environment) (opts : Options) (decl : Declaration) : Except KernelException Environment := do + let env ← addDecl env opts decl + compileDecl env opts decl + +def addDecl (decl : Declaration) : CoreM Unit := do + profileitM Exception "type checking" (← getOptions) do + withTraceNode `Kernel (fun _ => return m!"typechecking declaration") do + if !(← MonadLog.hasErrors) && decl.hasSorry then + logWarning "declaration uses 'sorry'" + match (← getEnv).addDecl (← getOptions) decl with + | .ok env => setEnv env + | .error ex => throwKernelException ex + +def addAndCompile (decl : Declaration) : CoreM Unit := do + addDecl decl + compileDecl decl + +end Lean diff --git a/src/Lean/Compiler/InitAttr.lean b/src/Lean/Compiler/InitAttr.lean index 4e723152fa75..4bc0dd9b9048 100644 --- a/src/Lean/Compiler/InitAttr.lean +++ b/src/Lean/Compiler/InitAttr.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ prelude +import Lean.AddDecl import Lean.Elab.InfoTree.Main namespace Lean diff --git a/src/Lean/Compiler/Old.lean b/src/Lean/Compiler/Old.lean index 2ae6cc23b365..ef87e3ba7ff5 100644 --- a/src/Lean/Compiler/Old.lean +++ b/src/Lean/Compiler/Old.lean @@ -67,9 +67,4 @@ opaque compileDecls (env : Environment) (opt : @& Options) (decls : @& List Name def compileDecl (env : Environment) (opt : @& Options) (decl : @& Declaration) : Except KernelException Environment := compileDecls env opt (Compiler.getDeclNamesForCodeGen decl) - -def addAndCompile (env : Environment) (opt : Options) (decl : Declaration) : Except KernelException Environment := do - let env ← addDecl env decl - compileDecl env opt decl - end Environment diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index 5664bd33196f..c92760026cdd 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -338,15 +338,6 @@ def mkArrow (d b : Expr) : CoreM Expr := /-- Iterated `mkArrow`, creates the expression `a₁ → a₂ → … → aₙ → b`. Also see `arrowDomainsN`. -/ def mkArrowN (ds : Array Expr) (e : Expr) : CoreM Expr := ds.foldrM mkArrow e -def addDecl (decl : Declaration) : CoreM Unit := do - profileitM Exception "type checking" (← getOptions) do - withTraceNode `Kernel (fun _ => return m!"typechecking declaration") do - if !(← MonadLog.hasErrors) && decl.hasSorry then - logWarning "declaration uses 'sorry'" - match (← getEnv).addDecl decl with - | Except.ok env => setEnv env - | Except.error ex => throwKernelException ex - private def supportedRecursors := #[``Empty.rec, ``False.rec, ``Eq.ndrec, ``Eq.rec, ``Eq.recOn, ``Eq.casesOn, ``False.casesOn, ``Empty.casesOn, ``And.rec, ``And.casesOn] @@ -400,10 +391,6 @@ def compileDecls (decls : List Name) : CoreM Unit := do | Except.error ex => throwKernelException ex -def addAndCompile (decl : Declaration) : CoreM Unit := do - addDecl decl; - compileDecl decl - def getDiag (opts : Options) : Bool := diagnostics.get opts diff --git a/src/Lean/Elab/BuiltinCommand.lean b/src/Lean/Elab/BuiltinCommand.lean index f04669e2242e..404ecab9bf08 100644 --- a/src/Lean/Elab/BuiltinCommand.lean +++ b/src/Lean/Elab/BuiltinCommand.lean @@ -123,7 +123,7 @@ private partial def elabChoiceAux (cmds : Array Syntax) (i : Nat) : CommandElabM n[1].forArgsM addUnivLevel @[builtin_command_elab «init_quot»] def elabInitQuot : CommandElab := fun _ => do - match (← getEnv).addDecl Declaration.quotDecl with + match (← getEnv).addDecl (← getOptions) Declaration.quotDecl with | Except.ok env => setEnv env | Except.error ex => throwError (ex.toMessageData (← getOptions)) diff --git a/src/Lean/Environment.lean b/src/Lean/Environment.lean index 1bb99cd02053..1ade563ad3a4 100644 --- a/src/Lean/Environment.lean +++ b/src/Lean/Environment.lean @@ -246,7 +246,7 @@ namespace Environment /-- Type check given declaration and add it to the environment -/ @[extern "lean_add_decl"] -opaque addDecl (env : Environment) (decl : @& Declaration) : Except KernelException Environment +opaque addDeclCore (env : Environment) (maxHeartbeats : USize) (decl : @& Declaration) : Except KernelException Environment end Environment diff --git a/src/Lean/Meta/Closure.lean b/src/Lean/Meta/Closure.lean index f50c36003029..75827f2839c9 100644 --- a/src/Lean/Meta/Closure.lean +++ b/src/Lean/Meta/Closure.lean @@ -6,6 +6,7 @@ Authors: Leonardo de Moura prelude import Lean.MetavarContext import Lean.Environment +import Lean.AddDecl import Lean.Util.FoldConsts import Lean.Meta.Basic import Lean.Meta.Check diff --git a/src/Lean/Meta/Constructions.lean b/src/Lean/Meta/Constructions.lean index 4eabaf48858a..2e4c9808d453 100644 --- a/src/Lean/Meta/Constructions.lean +++ b/src/Lean/Meta/Constructions.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Lean.AuxRecursor +import Lean.AddDecl import Lean.Meta.AppBuilder namespace Lean diff --git a/src/Lean/Meta/Eqns.lean b/src/Lean/Meta/Eqns.lean index c3ecc055c571..8e29d7a04cbc 100644 --- a/src/Lean/Meta/Eqns.lean +++ b/src/Lean/Meta/Eqns.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Lean.ReservedNameAction +import Lean.AddDecl import Lean.Meta.Basic import Lean.Meta.AppBuilder import Lean.Meta.Match.MatcherInfo diff --git a/src/Lean/Meta/Eval.lean b/src/Lean/Meta/Eval.lean index d4aa167213cd..8ad4784ec000 100644 --- a/src/Lean/Meta/Eval.lean +++ b/src/Lean/Meta/Eval.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Sebastian Ullrich, Leonardo de Moura -/ prelude +import Lean.AddDecl import Lean.Meta.Check namespace Lean.Meta diff --git a/src/Lean/Meta/SizeOf.lean b/src/Lean/Meta/SizeOf.lean index fa38b984cc1e..291301b8b426 100644 --- a/src/Lean/Meta/SizeOf.lean +++ b/src/Lean/Meta/SizeOf.lean @@ -5,6 +5,7 @@ Authors: Leonardo de Moura -/ prelude import Init.Data.List.BasicAux +import Lean.AddDecl import Lean.Meta.AppBuilder import Lean.Meta.Instances diff --git a/src/Lean/Meta/Tactic/AuxLemma.lean b/src/Lean/Meta/Tactic/AuxLemma.lean index eeb21b8c0dee..744ff2b90758 100644 --- a/src/Lean/Meta/Tactic/AuxLemma.lean +++ b/src/Lean/Meta/Tactic/AuxLemma.lean @@ -4,6 +4,7 @@ Released under Apache 2.0 license as described in the file LICENSE. Authors: Leonardo de Moura -/ prelude +import Lean.AddDecl import Lean.Meta.Basic namespace Lean.Meta diff --git a/src/Lean/Replay.lean b/src/Lean/Replay.lean index d44392191ec6..15bea755d9b7 100644 --- a/src/Lean/Replay.lean +++ b/src/Lean/Replay.lean @@ -5,6 +5,7 @@ Authors: Scott Morrison -/ prelude import Lean.CoreM +import Lean.AddDecl import Lean.Util.FoldConsts /-! @@ -56,7 +57,7 @@ def throwKernelException (ex : KernelException) : M Unit := do /-- Add a declaration, possibly throwing a `KernelException`. -/ def addDecl (d : Declaration) : M Unit := do - match (← get).env.addDecl d with + match (← get).env.addDecl {} d with | .ok env => modify fun s => { s with env := env } | .error ex => throwKernelException ex diff --git a/src/kernel/environment.cpp b/src/kernel/environment.cpp index d8316d5c9f3e..3e2d0be2ac30 100644 --- a/src/kernel/environment.cpp +++ b/src/kernel/environment.cpp @@ -292,7 +292,8 @@ environment environment::add(declaration const & d, bool check) const { lean_unreachable(); } -extern "C" LEAN_EXPORT object * lean_add_decl(object * env, object * decl) { +extern "C" LEAN_EXPORT object * lean_add_decl(object * env, size_t max_heartbeat, object * decl) { + scope_max_heartbeat s(max_heartbeat); return catch_kernel_exceptions([&]() { return environment(env).add(declaration(decl, true)); }); diff --git a/src/runtime/interrupt.cpp b/src/runtime/interrupt.cpp index 13c242582a02..6a0751ad5717 100644 --- a/src/runtime/interrupt.cpp +++ b/src/runtime/interrupt.cpp @@ -26,7 +26,7 @@ size_t get_max_heartbeat() { return g_max_heartbeat; } void set_max_heartbeat_thousands(unsigned max) { g_max_heartbeat = static_cast(max) * 1000; } scope_heartbeat::scope_heartbeat(size_t max):flet(g_heartbeat, max) {} -scope_max_heartbeat::scope_max_heartbeat(size_t max):flet(g_max_heartbeat, max) {} +LEAN_EXPORT scope_max_heartbeat::scope_max_heartbeat(size_t max):flet(g_max_heartbeat, max) {} // separate definition to allow breakpoint in debugger void throw_heartbeat_exception() { diff --git a/src/runtime/interrupt.h b/src/runtime/interrupt.h index ccc679319305..0f6b10cc5195 100644 --- a/src/runtime/interrupt.h +++ b/src/runtime/interrupt.h @@ -35,9 +35,9 @@ LEAN_EXPORT void set_max_heartbeat_thousands(unsigned max); LEAN_EXPORT size_t get_max_heartbeat(); /* Update the thread local max heartbeat */ -class scope_max_heartbeat : flet { +class LEAN_EXPORT scope_max_heartbeat : flet { public: - scope_max_heartbeat(size_t max); + LEAN_EXPORT scope_max_heartbeat(size_t max); }; LEAN_EXPORT void check_heartbeat(); diff --git a/tests/lean/partialIssue.lean b/tests/lean/partialIssue.lean index 55dc64c0d5cd..3684a969de5c 100644 --- a/tests/lean/partialIssue.lean +++ b/tests/lean/partialIssue.lean @@ -1,5 +1,5 @@ import Lean.CoreM - +import Lean.AddDecl #eval Lean.addDecl <| .mutualDefnDecl [{ name := `False_intro levelParams := [] diff --git a/tests/lean/run/kernel_maxheartbeats.lean b/tests/lean/run/kernel_maxheartbeats.lean new file mode 100644 index 000000000000..fa92ea48f9e5 --- /dev/null +++ b/tests/lean/run/kernel_maxheartbeats.lean @@ -0,0 +1,22 @@ +import Lean + +def ack : Nat → Nat → Nat + | 0, y => y+1 + | x+1, 0 => ack x 1 + | x+1, y+1 => ack x (ack (x+1) y) + +set_option maxHeartbeats 500 +open Lean Meta + +/-- +error: (kernel) deterministic timeout +-/ +#guard_msgs in +run_meta do + let type ← mkEq (← mkAppM ``ack #[mkNatLit 4, mkNatLit 4]) (mkNatLit 100000) + let value ← mkEqRefl (mkNatLit 100000) + addDecl <| .thmDecl { + name := `ack_4_4 + levelParams := [] + type, value + } diff --git a/tests/lean/run/structInstFast.lean b/tests/lean/run/structInstFast.lean index 38551a4aa0e5..ecc40bc43f99 100644 --- a/tests/lean/run/structInstFast.lean +++ b/tests/lean/run/structInstFast.lean @@ -92,7 +92,12 @@ def foo : String := "foo" deep_wide_struct_inst_with String foo 50 20 -/- Structure instances using the `with` pattern should be fast. Without #2478, this takes over 700 - heartbeats. -/ -set_option maxHeartbeats 200 in +/- +Structure instances using the `with` pattern should be fast. +Without #2478, this takes over 700 heartbeats in the elaborator. + +Remark: we are now propagating heartbeats to the kernel, and +had to increase it value to 1000 + -/ +set_option maxHeartbeats 1000 in example : a0 = a'0 := rfl From 368adaf84787c42823d2e4cdef7e3d1f7247817d Mon Sep 17 00:00:00 2001 From: Tobias Grosser Date: Fri, 10 May 2024 01:57:00 +0100 Subject: [PATCH 24/44] feat: add BitVec.[toInt_inj|toInt_ne] (#4075) Co-authored-by: Kim Morrison --- src/Init/Data/BitVec/Lemmas.lean | 6 ++++++ 1 file changed, 6 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 5a4fbcf1d9c7..a998800ea12a 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -245,6 +245,12 @@ theorem eq_of_toInt_eq {i j : BitVec n} : i.toInt = j.toInt → i = j := by have _jlt := j.isLt split <;> split <;> omega +theorem toInt_inj (x y : BitVec n) : x.toInt = y.toInt ↔ x = y := + Iff.intro eq_of_toInt_eq (congrArg BitVec.toInt) + +theorem toInt_ne (x y : BitVec n) : x.toInt ≠ y.toInt ↔ x ≠ y := by + rw [Ne, toInt_inj] + @[simp] theorem toNat_ofInt {n : Nat} (i : Int) : (BitVec.ofInt n i).toNat = (i % (2^n : Nat)).toNat := by unfold BitVec.ofInt From 3491c56c49e0abd6790acbb9c21b432ff8ebae3e Mon Sep 17 00:00:00 2001 From: Mario Carneiro Date: Fri, 10 May 2024 03:38:38 +0200 Subject: [PATCH 25/44] fix: segfault in old compiler due to noConfusion assumptions (#2903) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit This fixes #2901, a bug in the old compiler which causes a segfault. The issue is that when visiting `noConfusion` applications, it assumes that each constructor case has `nfields` arguments, e.g. `head1 = head2 -> tail1 = tail2 -> P` has two arguments because `List.cons` has 2 fields, but in fact propositional fields are skipped by the noConfusion type generator, so for example `Subtype.noConfusionType` is: ```lean @[reducible] protected def Subtype.noConfusionType.{u_1, u} : {α : Sort u} → {p : α → Prop} → Sort u_1 → Subtype p → Subtype p → Sort u_1 := fun {α} {p} P v1 v2 ↦ Subtype.casesOn v1 fun val property ↦ Subtype.casesOn v2 fun val_1 property ↦ (val = val_1 → P) → P ``` where `val = val_1 → P` only has the one argument even though `Subtype.mk` has two fields, presumably because it is useless to have an equality of propositions. Unfortunately there isn't any easy cache or getter to use here to get the number of non-propositional fields, so we just calculate it on the spot. --- src/library/compiler/lcnf.cpp | 19 ++++++++++++++++++- tests/lean/run/2901.lean | 14 ++++++++++++++ 2 files changed, 32 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/2901.lean diff --git a/src/library/compiler/lcnf.cpp b/src/library/compiler/lcnf.cpp index af2af4245272..9bcd88d4cf59 100644 --- a/src/library/compiler/lcnf.cpp +++ b/src/library/compiler/lcnf.cpp @@ -266,6 +266,23 @@ class to_lcnf_fn { return e; } + unsigned get_constructor_non_prop_nfields(name ctor, unsigned nparams) { + local_ctx lctx; + expr type = env().get(ctor).get_type(); + for (unsigned i = 0; i < nparams; i++) { + lean_assert(is_pi(type)); + expr local = lctx.mk_local_decl(ngen(), binding_name(type), binding_domain(type), binding_info(type)); + type = instantiate(binding_body(type), local); + } + unsigned nfields = 0; + while (is_pi(type)) { + if (!type_checker(m_st, lctx).is_prop(binding_domain(type))) nfields++; + expr local = lctx.mk_local_decl(ngen(), binding_name(type), binding_domain(type), binding_info(type)); + type = instantiate(binding_body(type), local); + } + return nfields; + } + expr visit_no_confusion(expr const & fn, buffer & args, bool root) { name const & no_confusion_name = const_name(fn); name const & I_name = no_confusion_name.get_prefix(); @@ -297,7 +314,7 @@ class to_lcnf_fn { lean_assert(args.size() >= basic_arity + 1); unsigned major_idx = basic_arity; expr major = args[major_idx]; - unsigned nfields = get_constructor_nfields(*lhs_constructor); + unsigned nfields = get_constructor_non_prop_nfields(*lhs_constructor, nparams); while (nfields > 0) { if (!is_lambda(major)) major = eta_expand(major, nfields); diff --git a/tests/lean/run/2901.lean b/tests/lean/run/2901.lean new file mode 100644 index 000000000000..ef51eccb7bfd --- /dev/null +++ b/tests/lean/run/2901.lean @@ -0,0 +1,14 @@ +def Vector (α : Type u) (n : Nat) := + { l : List α // l.length = n } + +inductive HVect : (n : Nat) -> (Vector (Type v) n) -> Type (v+1) where + | Nil : HVect 0 ⟨ [], simp ⟩ + | Cons : (t : Type v) -> (x : t) -> HVect n ⟨ts, p⟩ -> HVect (n+1) ⟨t::ts, by simp [p]⟩ + +def printHOK (v : HVect (n+1) ⟨String::ts, p'⟩) : String := + match v with + | HVect.Cons _ x _ => (x : String) + +def printHKO (v : HVect (n+1) ⟨String::ts, p'⟩) : String := + match v with + | HVect.Cons _ x _ => "Hi" From ca6437df717b550cbea3e110f3093c1250d5c5be Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Fri, 10 May 2024 01:12:19 -0400 Subject: [PATCH 26/44] fix: lake: TOML key order bug in `ppTable` (#4104) Fixes a bug in `Lake.Toml.ppTable` where root table keys could be printed after a subtable header. Closes #4099. --- src/lake/Lake/Toml/Data/Value.lean | 22 ++++++++++++------- .../tests/translateConfig/out.expected.toml | 3 +++ src/lake/tests/translateConfig/source.lean | 2 +- 3 files changed, 18 insertions(+), 9 deletions(-) diff --git a/src/lake/Lake/Toml/Data/Value.lean b/src/lake/Lake/Toml/Data/Value.lean index beb973d6905a..fb0618837d69 100644 --- a/src/lake/Lake/Toml/Data/Value.lean +++ b/src/lake/Lake/Toml/Data/Value.lean @@ -104,22 +104,28 @@ end instance : ToString Value := ⟨Value.toString⟩ def ppTable (t : Table) : String := - String.trimLeft <| t.items.foldl (init := "") fun s (k,v) => + let (ts, fs) := t.items.foldl (init := ("", "")) fun (ts, fs) (k,v) => match v with | .array _ vs => if vs.all (· matches .table ..) then - vs.foldl (init := s) fun s v => + let fs := vs.foldl (init := fs) fun s v => match v with | .table _ t => - let s := s ++ s!"\n[[{ppKey k}]]\n" - t.items.foldl (fun s (k,v) => appendKeyval s k v) s + let s := s ++ s!"[[{ppKey k}]]\n" + let s := t.items.foldl (fun s (k,v) => appendKeyval s k v) s + s.push '\n' | _ => unreachable! + (ts, fs) else - s.append s!"{ppKey k} = {ppInlineArray vs}\n" + (ts.append s!"{ppKey k} = {ppInlineArray vs}\n", fs) | .table _ t => - let s := s ++ s!"\n[{ppKey k}]\n" - t.items.foldl (fun s (k,v) => appendKeyval s k v) s - | _ => appendKeyval s k v + let fs := fs ++ s!"[{ppKey k}]\n" + let fs := t.items.foldl (fun s (k,v) => appendKeyval s k v) fs + (ts, fs.push '\n') + | _ => (appendKeyval ts k v, fs) + -- Ensures root table keys come before subtables + -- See https://github.com/leanprover/lean4/issues/4099 + (ts.push '\n' ++ fs).trimRight.push '\n' where appendKeyval s k v := s.append s!"{ppKey k} = {v}\n" diff --git a/src/lake/tests/translateConfig/out.expected.toml b/src/lake/tests/translateConfig/out.expected.toml index 89b532b4a7cb..6530167e9b1b 100644 --- a/src/lake/tests/translateConfig/out.expected.toml +++ b/src/lake/tests/translateConfig/out.expected.toml @@ -2,6 +2,9 @@ name = "test" testRunner = "b" defaultTargets = ["A"] +[leanOptions] +pp.unicode.fun = true + [[require]] name = "foo" path = "dir" diff --git a/src/lake/tests/translateConfig/source.lean b/src/lake/tests/translateConfig/source.lean index 83c74f58c3c9..00fd81f309fc 100644 --- a/src/lake/tests/translateConfig/source.lean +++ b/src/lake/tests/translateConfig/source.lean @@ -29,7 +29,7 @@ package test where buildArchive? := none preferReleaseBuild := false packagesDir := defaultPackagesDir - toLeanConfig := testLeanConfig + leanOptions := #[⟨`pp.unicode.fun, true⟩] require foo from "dir" with NameMap.empty.insert `foo "bar" require bar from git "https://example.com" @ "abc" / "sub" / "dir" From 39286862e3fa8752ea2cf0754b824a91fc5cd4b7 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Fri, 10 May 2024 08:45:21 +0200 Subject: [PATCH 27/44] feat: well-founded definitions irreducible by default (#4061) MIME-Version: 1.0 Content-Type: text/plain; charset=UTF-8 Content-Transfer-Encoding: 8bit we keep running into examples where working with well-founded recursion is slow because defeq checks (which are all over the place, including failing ones that are back-tracked) unfold well-founded definitions. The definition of a function defined by well-founded recursion should be an implementation detail that should only be peeked inside by the equation generator and the functional induction generator. We now mark the mutual recursive function as irreducible (if the user did not set a flag explicitly), and use `withAtLeastTransparency .all` when producing the equations. Proofs can be fixed by using rewriting, or – a bit blunt, but nice for adjusting existing proofs – using `unseal` (a.k.a. `attribute [local semireducible]`). Mathlib performance does not change a whole lot: http://speed.lean-fro.org/mathlib4/compare/08b82265-75db-4a28-b12b-08751b9ad04a/to/16f46d5e-28b1-41c4-a107-a6f6594841f8 Build instructions -0.126 %, four modules with significant instructions decrease. To reduce impact, these definitions were changed: * `Nat.mod`, to make `1 % n` reduce definitionally, so that `1` as a `Fin 2` literal works nicely. Theorems with larger `Fin` literals tend to need a `unseal Nat.modCore` https://github.com/leanprover/lean4/pull/4098 * `List.ofFn` rewritten to be structurally recursive and not go via `Array.ofFn`: https://github.com/leanprover-community/batteries/pull/784 Alternative designs explored were * Making `WellFounded.fix` irreducible. One benefit is that recursive functions with equal definitions (possibly after instantiating fixed parameters) are defeq; this is used in mathlib to relate [`OrdinalApprox.gfpApprox`](https://leanprover-community.github.io/mathlib4_docs/Mathlib/SetTheory/Ordinal/FixedPointApproximants.html#OrdinalApprox.gfpApprox) with `.lfpApprox`. But the downside is that one cannot use `unseal` in a targeted way, being explicit in which recursive function needs to be reducible here. And in cases where Lean does unwanted unfolding, we’d still unfold the recursive definition once to expose `WellFounded.fix`, leading to large terms for often no good reason. * Defining `WellFounded.fix` to unroll defintionally once before hitting a irreducible `WellFounded.fixF`. This was explored in #4002. It shares most of the ups and downs with the previous variant, with the additional neat benefit that function calls that do not lead to recursive cases (e.g. a `[]` base case) reduce nicely. This means that the majority of existing `rfl` proofs continue to work. Issue #4051, which demonstrates how badly things can go if wf recursive functions can be unrolled, showed that making the recursive function irreducible there leads to noticeably faster elaboration than making `WellFounded.fix` irreducible; this is good evidence that the present PR is the way to go. This fixes https://github.com/leanprover/lean4/issues/3988 --------- Co-authored-by: Leonardo de Moura --- RELEASES.md | 13 +- src/Init/Data/Fin/Lemmas.lean | 6 + src/Init/Data/Int/DivModLemmas.lean | 7 + src/Init/Data/Nat/Bitwise/Lemmas.lean | 5 +- src/Init/Data/Nat/Gcd.lean | 10 +- src/Init/Data/Nat/Lemmas.lean | 6 +- src/Lean/Elab/PreDefinition/WF/Eqns.lean | 2 +- src/Lean/Elab/PreDefinition/WF/Main.lean | 17 ++- src/Lean/ReducibilityAttrs.lean | 5 + stage0/src/stdlib_flags.h | 2 + tests/lean/run/1921.lean | 2 +- tests/lean/run/2389.lean | 4 +- tests/lean/run/ack.lean | 1 + tests/lean/run/casesRec.lean | 2 +- tests/lean/run/defaultEliminator.lean | 2 +- tests/lean/run/lazylistThunk.lean | 2 +- tests/lean/run/splitIssue.lean | 16 +-- tests/lean/run/splitList.lean | 32 ++--- tests/lean/run/wfirred.lean | 139 ++++++++++++++++++++ tests/lean/wfrecUnusedLet.lean.expected.out | 2 +- 20 files changed, 228 insertions(+), 47 deletions(-) create mode 100644 tests/lean/run/wfirred.lean diff --git a/RELEASES.md b/RELEASES.md index 940cc2de9b6f..32378cc2a64b 100644 --- a/RELEASES.md +++ b/RELEASES.md @@ -11,7 +11,18 @@ of each version. v4.9.0 (development in progress) --------- -v4.8.0 +* Functions defined by well-founded recursion are now marked as + `@[irreducible]`, which should prevent expensive and often unfruitful + unfolding of such definitions. + + Existing proofs that hold by definitional equality (e.g. `rfl`) can be + rewritten to explictly unfold the function definition (using `simp`, + `unfold`, `rw`), or the recursive function can be temporariliy made + semireducible (using `unseal f in` before the command) or the function + definition itself can be marked as `@[semireducible]` to get the previous + behavor. + +v4.8.0 --------- * **Executables configured with `supportInterpreter := true` on Windows should now be run via `lake exe` to function properly.** diff --git a/src/Init/Data/Fin/Lemmas.lean b/src/Init/Data/Fin/Lemmas.lean index 3f9fc0f431c2..1a0f5b9113e3 100644 --- a/src/Init/Data/Fin/Lemmas.lean +++ b/src/Init/Data/Fin/Lemmas.lean @@ -11,6 +11,9 @@ import Init.ByCases import Init.Conv import Init.Omega +-- Remove after the next stage0 update +set_option allowUnsafeReducibility true + namespace Fin /-- If you actually have an element of `Fin n`, then the `n` is always positive -/ @@ -205,6 +208,7 @@ theorem val_add_one {n : Nat} (i : Fin (n + 1)) : | .inl h => cases Fin.eq_of_val_eq h; simp | .inr h => simpa [Fin.ne_of_lt h] using val_add_one_of_lt h +unseal Nat.modCore in @[simp] theorem val_two {n : Nat} : (2 : Fin (n + 3)).val = 2 := rfl theorem add_one_pos (i : Fin (n + 1)) (h : i < Fin.last n) : (0 : Fin (n + 1)) < i + 1 := by @@ -239,6 +243,7 @@ theorem succ_ne_zero {n} : ∀ k : Fin n, Fin.succ k ≠ 0 @[simp] theorem succ_zero_eq_one : Fin.succ (0 : Fin (n + 1)) = 1 := rfl +unseal Nat.modCore in /-- Version of `succ_one_eq_two` to be used by `dsimp` -/ @[simp] theorem succ_one_eq_two : Fin.succ (1 : Fin (n + 2)) = 2 := rfl @@ -390,6 +395,7 @@ theorem castSucc_lt_last (a : Fin n) : castSucc a < last n := a.is_lt @[simp] theorem castSucc_zero : castSucc (0 : Fin (n + 1)) = 0 := rfl +unseal Nat.modCore in @[simp] theorem castSucc_one {n : Nat} : castSucc (1 : Fin (n + 2)) = 1 := rfl /-- `castSucc i` is positive when `i` is positive -/ diff --git a/src/Init/Data/Int/DivModLemmas.lean b/src/Init/Data/Int/DivModLemmas.lean index d532d40a4272..a6a33b293f0e 100644 --- a/src/Init/Data/Int/DivModLemmas.lean +++ b/src/Init/Data/Int/DivModLemmas.lean @@ -14,6 +14,8 @@ import Init.RCases # Lemmas about integer division needed to bootstrap `omega`. -/ +-- Remove after the next stage0 update +set_option allowUnsafeReducibility true open Nat (succ) @@ -142,12 +144,14 @@ theorem eq_one_of_mul_eq_one_left {a b : Int} (H : 0 ≤ b) (H' : a * b = 1) : b | ofNat _ => show ofNat _ = _ by simp | -[_+1] => show -ofNat _ = _ by simp +unseal Nat.div in @[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 +unseal Nat.div in @[simp] protected theorem fdiv_zero : ∀ a : Int, fdiv a 0 = 0 | 0 => rfl | succ _ => rfl @@ -765,11 +769,13 @@ theorem ediv_eq_ediv_of_mul_eq_mul {a b c d : Int} | (n:Nat) => congrArg ofNat (Nat.div_one _) | -[n+1] => by simp [Int.div, neg_ofNat_succ]; rfl +unseal Nat.div in @[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 +unseal Nat.div in @[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 @@ -938,6 +944,7 @@ 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 _ +unseal Nat.div in theorem fdiv_nonpos : ∀ {a b : Int}, 0 ≤ a → b ≤ 0 → a.fdiv b ≤ 0 | 0, 0, _, _ | 0, -[_+1], _, _ | succ _, 0, _, _ | succ _, -[_+1], _, _ => ⟨_⟩ diff --git a/src/Init/Data/Nat/Bitwise/Lemmas.lean b/src/Init/Data/Nat/Bitwise/Lemmas.lean index c59bff8945ec..53467f29827a 100644 --- a/src/Init/Data/Nat/Bitwise/Lemmas.lean +++ b/src/Init/Data/Nat/Bitwise/Lemmas.lean @@ -50,7 +50,10 @@ noncomputable def div2Induction {motive : Nat → Sort u} apply hyp exact Nat.div_lt_self n_pos (Nat.le_refl _) -@[simp] theorem zero_and (x : Nat) : 0 &&& x = 0 := by rfl +@[simp] theorem zero_and (x : Nat) : 0 &&& x = 0 := by + simp only [HAnd.hAnd, AndOp.and, land] + unfold bitwise + simp @[simp] theorem and_zero (x : Nat) : x &&& 0 = 0 := by simp only [HAnd.hAnd, AndOp.and, land] diff --git a/src/Init/Data/Nat/Gcd.lean b/src/Init/Data/Nat/Gcd.lean index c5a633e586a3..78961e04e034 100644 --- a/src/Init/Data/Nat/Gcd.lean +++ b/src/Init/Data/Nat/Gcd.lean @@ -37,11 +37,11 @@ def gcd (m n : @& Nat) : Nat := 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 +@[simp] theorem gcd_zero_left (y : Nat) : gcd 0 y = y := by + rw [gcd]; rfl -theorem gcd_succ (x y : Nat) : gcd (succ x) y = gcd (y % succ x) (succ x) := - rfl +theorem gcd_succ (x y : Nat) : gcd (succ x) y = gcd (y % succ x) (succ x) := by + rw [gcd]; rfl @[simp] theorem gcd_one_left (n : Nat) : gcd 1 n = 1 := by rw [gcd_succ, mod_one] @@ -64,7 +64,7 @@ instance : Std.IdempotentOp gcd := ⟨gcd_self⟩ theorem gcd_rec (m n : Nat) : gcd m n = gcd (n % m) m := match m with - | 0 => by have := (mod_zero n).symm; rwa [gcd_zero_right] + | 0 => by have := (mod_zero n).symm; rwa [gcd, gcd_zero_right] | _ + 1 => by simp [gcd_succ] @[elab_as_elim] theorem gcd.induction {P : Nat → Nat → Prop} (m n : Nat) diff --git a/src/Init/Data/Nat/Lemmas.lean b/src/Init/Data/Nat/Lemmas.lean index 50124336f25d..af9f3104b1a5 100644 --- a/src/Init/Data/Nat/Lemmas.lean +++ b/src/Init/Data/Nat/Lemmas.lean @@ -677,6 +677,10 @@ protected theorem pow_lt_pow_iff_right {a n m : Nat} (h : 1 < a) : /-! ### log2 -/ +@[simp] +theorem log2_zero : Nat.log2 0 = 0 := by + simp [Nat.log2] + theorem le_log2 (h : n ≠ 0) : k ≤ n.log2 ↔ 2 ^ k ≤ n := by match k with | 0 => simp [show 1 ≤ n from Nat.pos_of_ne_zero h] @@ -697,7 +701,7 @@ theorem log2_self_le (h : n ≠ 0) : 2 ^ n.log2 ≤ n := (le_log2 h).1 (Nat.le_r theorem lt_log2_self : n < 2 ^ (n.log2 + 1) := match n with - | 0 => Nat.zero_lt_two + | 0 => by simp | n+1 => (log2_lt n.succ_ne_zero).1 (Nat.le_refl _) /-! ### dvd -/ diff --git a/src/Lean/Elab/PreDefinition/WF/Eqns.lean b/src/Lean/Elab/PreDefinition/WF/Eqns.lean index 8b1c6ae68928..41d69f15b846 100644 --- a/src/Lean/Elab/PreDefinition/WF/Eqns.lean +++ b/src/Lean/Elab/PreDefinition/WF/Eqns.lean @@ -81,7 +81,7 @@ private partial def mkProof (declName : Name) (type : Expr) : MetaM Expr := do let (_, mvarId) ← main.mvarId!.intros let rec go (mvarId : MVarId) : MetaM Unit := do trace[Elab.definition.wf.eqns] "step\n{MessageData.ofGoal mvarId}" - if (← tryURefl mvarId) then + if ← withAtLeastTransparency .all (tryURefl mvarId) then return () else if (← tryContradiction mvarId) then return () diff --git a/src/Lean/Elab/PreDefinition/WF/Main.lean b/src/Lean/Elab/PreDefinition/WF/Main.lean index f05f560bb71e..97a7560e8413 100644 --- a/src/Lean/Elab/PreDefinition/WF/Main.lean +++ b/src/Lean/Elab/PreDefinition/WF/Main.lean @@ -132,12 +132,15 @@ def wfRecursion (preDefs : Array PreDefinition) : TermElabM Unit := do return { unaryPreDef with value } trace[Elab.definition.wf] ">> {preDefNonRec.declName} :=\n{preDefNonRec.value}" let preDefs ← preDefs.mapM fun d => eraseRecAppSyntax d - if (← isOnlyOneUnaryDef preDefs fixedPrefixSize) then - addNonRec preDefNonRec (applyAttrAfterCompilation := false) - else - withEnableInfoTree false do + -- Do not complain if the user sets @[semireducible], which usually is a noop, + -- we recognize that below and then do not set @[irreducible] + withOptions (allowUnsafeReducibility.set · true) do + if (← isOnlyOneUnaryDef preDefs fixedPrefixSize) then addNonRec preDefNonRec (applyAttrAfterCompilation := false) - addNonRecPreDefs fixedPrefixSize argsPacker preDefs preDefNonRec + else + withEnableInfoTree false do + addNonRec preDefNonRec (applyAttrAfterCompilation := false) + addNonRecPreDefs fixedPrefixSize argsPacker preDefs preDefNonRec -- We create the `_unsafe_rec` before we abstract nested proofs. -- Reason: the nested proofs may be referring to the _unsafe_rec. addAndCompilePartialRec preDefs @@ -146,6 +149,10 @@ def wfRecursion (preDefs : Array PreDefinition) : TermElabM Unit := do for preDef in preDefs do markAsRecursive preDef.declName applyAttributesOf #[preDef] AttributeApplicationTime.afterCompilation + -- Unless the user asks for something else, mark the definition as irreducible + unless preDef.modifiers.attrs.any fun a => + a.name = `semireducible || a.name = `reducible || a.name = `semireducible do + setIrreducibleAttribute preDef.declName builtin_initialize registerTraceClass `Elab.definition.wf diff --git a/src/Lean/ReducibilityAttrs.lean b/src/Lean/ReducibilityAttrs.lean index f843f0e4fc27..2dac2267a2a5 100644 --- a/src/Lean/ReducibilityAttrs.lean +++ b/src/Lean/ReducibilityAttrs.lean @@ -184,4 +184,9 @@ def isIrreducible [Monad m] [MonadEnv m] (declName : Name) : m Bool := do | .irreducible => return true | _ => return false +/-- Set the given declaration as `[irreducible]` -/ +def setIrreducibleAttribute [Monad m] [MonadEnv m] (declName : Name) : m Unit := do + setReducibilityStatus declName ReducibilityStatus.irreducible + + end Lean diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 0699845ba452..158ed2fe6013 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -1,5 +1,7 @@ #include "util/options.h" +// please auto update stage0 + namespace lean { options get_default_options() { options opts; diff --git a/tests/lean/run/1921.lean b/tests/lean/run/1921.lean index bdde9e9c3d29..39cb99298764 100644 --- a/tests/lean/run/1921.lean +++ b/tests/lean/run/1921.lean @@ -6,4 +6,4 @@ def f := #[true].any id 0 USize.size -- `native_decide` used to prove `false` here, due to a bug in `Array.anyMUnsafe`. example : f = true := by native_decide -example : f = true := by simp (config := { decide := true }) [f, Array.any, Array.anyM] +example : f = true := by simp (config := { decide := true }) [f, Array.any, Array.anyM, Array.anyM.loop] diff --git a/tests/lean/run/2389.lean b/tests/lean/run/2389.lean index 6923e6c5e612..c1d6cc1df8a3 100644 --- a/tests/lean/run/2389.lean +++ b/tests/lean/run/2389.lean @@ -26,7 +26,9 @@ def onlyZeros : Tree → Prop | .node [] => True | .node (x::s) => onlyZeros x ∧ onlyZeros (.node s) -/-- Pattern-matching on `OnlyZeros` works despite `below` and `brecOn` not being generated. -/ +unseal onlyZeros in +/-- Pattern-matching on `OnlyZeros` works despite `below` and `brecOn` not being generated +if we make `onlyZeros` semireducible-/ def toFixPoint : OnlyZeros t → onlyZeros t | .leaf => rfl | .node [] _ => True.intro diff --git a/tests/lean/run/ack.lean b/tests/lean/run/ack.lean index 41fd2ed6ea9a..54c8166c8e40 100644 --- a/tests/lean/run/ack.lean +++ b/tests/lean/run/ack.lean @@ -28,6 +28,7 @@ info: [reduction] unfolded declarations (max: 1725, num: 4): Acc.rec ↦ 754use `set_option diagnostics.threshold ` to control threshold for reporting counters -/ #guard_msgs in +unseal ack in set_option diagnostics.threshold 500 in set_option diagnostics true in theorem ex : ack 3 2 = 29 := diff --git a/tests/lean/run/casesRec.lean b/tests/lean/run/casesRec.lean index 2adf9e10e199..a8f4c639a291 100644 --- a/tests/lean/run/casesRec.lean +++ b/tests/lean/run/casesRec.lean @@ -8,7 +8,7 @@ def f (x : Nat) : Nat := by #eval f 10 -example : f x.succ = 2 * f x := rfl +example : f x.succ = 2 * f x := by rw [f]; rfl end Ex1 namespace Ex2 diff --git a/tests/lean/run/defaultEliminator.lean b/tests/lean/run/defaultEliminator.lean index 588487931503..bc9c02f25a92 100644 --- a/tests/lean/run/defaultEliminator.lean +++ b/tests/lean/run/defaultEliminator.lean @@ -22,7 +22,7 @@ termination_by (x, y) example (x y : Nat) : f x y > 0 := by induction x, y with - | zero_zero => decide + | zero_zero => simp [f] | succ_zero x ih => simp [f, ih] | zero_succ y ih => simp [f, ih] | succ_succ x y ih => simp [f, ih] diff --git a/tests/lean/run/lazylistThunk.lean b/tests/lean/run/lazylistThunk.lean index fc3f451ffbca..785481da4fe3 100644 --- a/tests/lean/run/lazylistThunk.lean +++ b/tests/lean/run/lazylistThunk.lean @@ -25,7 +25,7 @@ in the list, ignoring delays theorem length_toList (l : LazyList α) : l.toList.length = l.length := by match l with - | nil => rfl + | nil => simp [length_toList] | cons a as => simp [length_toList as] | delayed as => simp [length_toList as.get] diff --git a/tests/lean/run/splitIssue.lean b/tests/lean/run/splitIssue.lean index b555f3541607..fa0ee1dca109 100644 --- a/tests/lean/run/splitIssue.lean +++ b/tests/lean/run/splitIssue.lean @@ -15,23 +15,21 @@ termination_by l => l.length decreasing_by all_goals sorry -theorem len_nil : len ([] : List α) = 0 := by - simp [len] - --- The `simp [len]` above generated the following equation theorems for len +-- The equational theorems are #check @len.eq_1 #check @len.eq_2 #check @len.eq_3 -- It is conditional, and may be tricky to use. +#check @len.eq_def + +theorem len_nil : len ([] : List α) = 0 := by + simp [len] theorem len_1 (a : α) : len [a] = 1 := by simp [len] theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by conv => lhs; unfold len - rfl - --- The `unfold` tactic above generated the following theorem -#check @len.eq_def + cases bs <;> simp [splitList, len_1] theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with @@ -41,7 +39,7 @@ theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by theorem listlen : ∀ l : List α, l.length = len l := by intro l induction l with - | nil => rfl + | nil => simp [len] | cons h t ih => simp [List.length, len_cons, ih] rw [Nat.add_comm] diff --git a/tests/lean/run/splitList.lean b/tests/lean/run/splitList.lean index 26f22f3a7562..5d0229eb27e2 100644 --- a/tests/lean/run/splitList.lean +++ b/tests/lean/run/splitList.lean @@ -34,23 +34,21 @@ def len : List α → Nat len fst + len snd termination_by xs => xs.length -theorem len_nil : len ([] : List α) = 0 := by - simp [len] --- The `simp [len]` above generated the following equation theorems for len +-- The equational theorems are #check @len.eq_1 #check @len.eq_2 #check @len.eq_3 +#check @len.eq_def + +theorem len_nil : len ([] : List α) = 0 := by + simp [len] theorem len_1 (a : α) : len [a] = 1 := by simp [len] theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by - conv => lhs; unfold len - rfl - --- The `unfold` tactic above generated the following theorem -#check @len.eq_def + simp [len, splitList] theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with @@ -60,7 +58,7 @@ theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by theorem listlen : ∀ l : List α, l.length = len l := by intro l induction l with - | nil => rfl + | nil => simp [len_nil] | cons h t ih => simp [List.length, len_cons, ih] rw [Nat.add_comm] @@ -85,23 +83,21 @@ decreasing_by subst h₂ simp_arith [eq_of_heq h₃] at this |- ; simp [this] -theorem len_nil : len ([] : List α) = 0 := by - simp [len] - --- The `simp [len]` above generated the following equation theorems for len +-- The equational theorems are #check @len.eq_1 #check @len.eq_2 #check @len.eq_3 +#check @len.eq_def + +theorem len_nil : len ([] : List α) = 0 := by + simp [len] theorem len_1 (a : α) : len [a] = 1 := by simp [len] theorem len_2 (a b : α) (bs : List α) : len (a::b::bs) = 1 + len (b::bs) := by conv => lhs; unfold len - rfl - --- The `unfold` tactic above generated the following theorem -#check @len.eq_def + simp [len, splitList] theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by cases as with @@ -111,7 +107,7 @@ theorem len_cons (a : α) (as : List α) : len (a::as) = 1 + len as := by theorem listlen : ∀ l : List α, l.length = len l := by intro l induction l with - | nil => rfl + | nil => simp [len_nil] | cons h t ih => simp [List.length, len_cons, ih] rw [Nat.add_comm] diff --git a/tests/lean/run/wfirred.lean b/tests/lean/run/wfirred.lean new file mode 100644 index 000000000000..6916071e95d7 --- /dev/null +++ b/tests/lean/run/wfirred.lean @@ -0,0 +1,139 @@ +/-! +Tests that definitions by well-founded recursion are irreducible. +-/ + +def foo : Nat → Nat + | 0 => 0 + | n+1 => foo n +termination_by n => n + +/-- +error: type mismatch + rfl +has type + foo 0 = foo 0 : Prop +but is expected to have type + foo 0 = 0 : Prop +-/ +#guard_msgs in +example : foo 0 = 0 := rfl + +/-- +error: type mismatch + rfl +has type + foo (n + 1) = foo (n + 1) : Prop +but is expected to have type + foo (n + 1) = foo n : Prop +-/ +#guard_msgs in +example : foo (n+1) = foo n := rfl + +-- This succeeding is a bug or misfeature in the rfl tactic, using the kernel defeq check +#guard_msgs in +example : foo 0 = 0 := by rfl + +-- It only works on closed terms: +/-- +error: The rfl tactic failed. Possible reasons: +- The goal is not a reflexive relation (neither `=` nor a relation with a @[refl] lemma). +- The arguments of the relation are not equal. +Try using the reflexivitiy lemma for your relation explicitly, e.g. `exact Eq.rfl`. +n : Nat +⊢ foo (n + 1) = foo n +-/ +#guard_msgs in +example : foo (n+1) = foo n := by rfl + +section Unsealed + +unseal foo + +example : foo 0 = 0 := rfl +example : foo 0 = 0 := by rfl + +example : foo (n+1) = foo n := rfl +example : foo (n+1) = foo n := by rfl + +end Unsealed + +--should be sealed again here + +/-- +error: type mismatch + rfl +has type + foo 0 = foo 0 : Prop +but is expected to have type + foo 0 = 0 : Prop +-/ +#guard_msgs in +example : foo 0 = 0 := rfl + + +def bar : Nat → Nat + | 0 => 0 + | n+1 => bar n +termination_by n => n + +-- Once unsealed, the full internals are visible. This allows one to prove, for example + +/-- +error: type mismatch + rfl +has type + foo = foo : Prop +but is expected to have type + foo = bar : Prop +-/ +#guard_msgs in +example : foo = bar := rfl + + +unseal foo bar in +example : foo = bar := rfl + + +-- Attributes on the definition take precedence +@[semireducible] def baz : Nat → Nat + | 0 => 0 + | n+1 => baz n +termination_by n => n + +example : baz 0 = 0 := rfl + +seal baz in +/-- +error: type mismatch + rfl +has type + baz 0 = baz 0 : Prop +but is expected to have type + baz 0 = 0 : Prop +-/ +#guard_msgs in +example : baz 0 = 0 := rfl + +example : baz 0 = 0 := rfl + +@[reducible] def quux : Nat → Nat + | 0 => 0 + | n+1 => quux n +termination_by n => n + +example : quux 0 = 0 := rfl + +set_option allowUnsafeReducibility true in +seal quux in +/-- +error: type mismatch + rfl +has type + quux 0 = quux 0 : Prop +but is expected to have type + quux 0 = 0 : Prop +-/ +#guard_msgs in +example : quux 0 = 0 := rfl + +example : quux 0 = 0 := rfl diff --git a/tests/lean/wfrecUnusedLet.lean.expected.out b/tests/lean/wfrecUnusedLet.lean.expected.out index 7fbd9db3dc08..a9c7f8c558a6 100644 --- a/tests/lean/wfrecUnusedLet.lean.expected.out +++ b/tests/lean/wfrecUnusedLet.lean.expected.out @@ -1,4 +1,4 @@ -def f : Nat → Nat := +@[irreducible] def f : Nat → Nat := f.proof_1.fix fun n a => if h : n = 0 then 1 else From dcdc3db3d441791bd35ee06489f5b9a8094cd64c Mon Sep 17 00:00:00 2001 From: Lean stage0 autoupdater <> Date: Fri, 10 May 2024 07:39:47 +0000 Subject: [PATCH 28/44] chore: update stage0 --- stage0/src/kernel/environment.cpp | 3 +- stage0/src/library/compiler/lcnf.cpp | 19 +- stage0/src/runtime/interrupt.cpp | 2 +- stage0/src/runtime/interrupt.h | 4 +- stage0/src/stdlib_flags.h | 2 - stage0/stdlib/Init/Control/Lawful/Basic.c | 32 +- stage0/stdlib/Init/GetElem.c | 8 +- stage0/stdlib/Init/System/IO.c | 778 +- stage0/stdlib/Lean.c | 6 +- stage0/stdlib/Lean/AddDecl.c | 2581 +++++ stage0/stdlib/Lean/Attributes.c | 194 +- stage0/stdlib/Lean/Class.c | 4 +- stage0/stdlib/Lean/Compiler/CSimpAttr.c | 6 +- stage0/stdlib/Lean/Compiler/IR/ResetReuse.c | 2571 ++--- stage0/stdlib/Lean/Compiler/InitAttr.c | 102 +- stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c | 56 +- stage0/stdlib/Lean/Compiler/LCNF/Main.c | 8 +- stage0/stdlib/Lean/Compiler/LCNF/Passes.c | 26 +- stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c | 70 +- stage0/stdlib/Lean/Compiler/Main.c | 8 +- stage0/stdlib/Lean/Compiler/Old.c | 48 - stage0/stdlib/Lean/CoreM.c | 4987 +++------ stage0/stdlib/Lean/Data/Options.c | 14 +- stage0/stdlib/Lean/Elab/App.c | 4 +- stage0/stdlib/Lean/Elab/Binders.c | 1132 +-- stage0/stdlib/Lean/Elab/BuiltinCommand.c | 1079 +- stage0/stdlib/Lean/Elab/BuiltinNotation.c | 1020 +- stage0/stdlib/Lean/Elab/BuiltinTerm.c | 4 +- stage0/stdlib/Lean/Elab/CheckTactic.c | 16 +- stage0/stdlib/Lean/Elab/Command.c | 517 +- stage0/stdlib/Lean/Elab/Declaration.c | 349 +- stage0/stdlib/Lean/Elab/Deriving/BEq.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/Basic.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/DecEq.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/FromToJson.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/Hashable.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/Inhabited.c | 6 +- stage0/stdlib/Lean/Elab/Deriving/Nonempty.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/Ord.c | 4 +- stage0/stdlib/Lean/Elab/Deriving/Repr.c | 923 +- stage0/stdlib/Lean/Elab/Deriving/Util.c | 4 +- stage0/stdlib/Lean/Elab/Eval.c | 167 +- stage0/stdlib/Lean/Elab/Extra.c | 2226 ++-- stage0/stdlib/Lean/Elab/Inductive.c | 658 +- stage0/stdlib/Lean/Elab/InheritDoc.c | 4 +- stage0/stdlib/Lean/Elab/Match.c | 1651 +-- stage0/stdlib/Lean/Elab/MutualDef.c | 192 +- .../stdlib/Lean/Elab/PreDefinition/WF/Eqns.c | 1337 ++- .../stdlib/Lean/Elab/PreDefinition/WF/Main.c | 4671 +++++---- .../PreDefinition/WF/TerminationArgument.c | 6 +- .../Elab/PreDefinition/WF/TerminationHint.c | 14 +- stage0/stdlib/Lean/Elab/Print.c | 8 +- stage0/stdlib/Lean/Elab/StructInst.c | 10 +- stage0/stdlib/Lean/Elab/Structure.c | 2226 ++-- stage0/stdlib/Lean/Elab/SyntheticMVars.c | 4872 +++++---- .../stdlib/Lean/Elab/Tactic/BuiltinTactic.c | 6 +- stage0/stdlib/Lean/Elab/Tactic/Conv/Rewrite.c | 11 +- stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c | 353 +- stage0/stdlib/Lean/Elab/Tactic/Guard.c | 460 +- stage0/stdlib/Lean/Elab/Tactic/Induction.c | 199 +- .../stdlib/Lean/Elab/Tactic/LibrarySearch.c | 1 + stage0/stdlib/Lean/Elab/Tactic/Meta.c | 21 +- stage0/stdlib/Lean/Elab/Tactic/NormCast.c | 195 +- .../stdlib/Lean/Elab/Tactic/Omega/Frontend.c | 242 +- stage0/stdlib/Lean/Elab/Tactic/RCases.c | 8955 +++++++++-------- stage0/stdlib/Lean/Elab/Tactic/Rewrite.c | 320 +- stage0/stdlib/Lean/Elab/Tactic/Rewrites.c | 3 + stage0/stdlib/Lean/Elab/Tactic/ShowTerm.c | 151 +- stage0/stdlib/Lean/Elab/Tactic/Simp.c | 467 +- stage0/stdlib/Lean/Elab/Tactic/Simproc.c | 87 +- stage0/stdlib/Lean/Elab/Tactic/SolveByElim.c | 253 +- stage0/stdlib/Lean/Elab/Term.c | 1120 +-- stage0/stdlib/Lean/Elab/Util.c | 216 +- stage0/stdlib/Lean/Environment.c | 526 +- stage0/stdlib/Lean/KeyedDeclsAttribute.c | 12 +- stage0/stdlib/Lean/LabelAttribute.c | 16 +- stage0/stdlib/Lean/Language/Basic.c | 340 +- stage0/stdlib/Lean/Language/Lean.c | 1967 ++-- stage0/stdlib/Lean/Linter/MissingDocs.c | 6 +- stage0/stdlib/Lean/Linter/UnusedVariables.c | 6 +- stage0/stdlib/Lean/Meta/Closure.c | 6 +- stage0/stdlib/Lean/Meta/CoeAttr.c | 4 +- stage0/stdlib/Lean/Meta/Constructions.c | 12 +- stage0/stdlib/Lean/Meta/Eqns.c | 6 +- stage0/stdlib/Lean/Meta/Eval.c | 6 +- stage0/stdlib/Lean/Meta/Instances.c | 4 +- stage0/stdlib/Lean/Meta/LazyDiscrTree.c | 948 +- stage0/stdlib/Lean/Meta/SizeOf.c | 8 +- stage0/stdlib/Lean/Meta/SynthInstance.c | 345 +- stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c | 12 +- .../stdlib/Lean/Meta/Tactic/LibrarySearch.c | 21 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c | 132 +- .../Lean/Meta/Tactic/Simp/SimpTheorems.c | 16 +- stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c | 606 +- stage0/stdlib/Lean/Meta/Transform.c | 8 +- stage0/stdlib/Lean/Parser/Extension.c | 462 +- stage0/stdlib/Lean/PrettyPrinter/Formatter.c | 4 +- .../stdlib/Lean/PrettyPrinter/Parenthesizer.c | 4 +- stage0/stdlib/Lean/ReducibilityAttrs.c | 29 + stage0/stdlib/Lean/Replay.c | 235 +- stage0/stdlib/Lean/ReservedNameAction.c | 992 +- stage0/stdlib/Lean/ScopedEnvExtension.c | 14 +- stage0/stdlib/Lean/Server/CodeActions/Attr.c | 8 +- stage0/stdlib/Lean/Server/CodeActions/Basic.c | 6 +- stage0/stdlib/Lean/Server/FileWorker.c | 1053 +- stage0/stdlib/Lean/Server/Rpc/Deriving.c | 4 +- .../stdlib/Lean/Server/Rpc/RequestHandling.c | 4 +- stage0/stdlib/Lean/Util/Heartbeats.c | 952 +- stage0/stdlib/Lean/Widget/UserWidget.c | 6 +- 109 files changed, 31409 insertions(+), 25084 deletions(-) create mode 100644 stage0/stdlib/Lean/AddDecl.c diff --git a/stage0/src/kernel/environment.cpp b/stage0/src/kernel/environment.cpp index d8316d5c9f3e..3e2d0be2ac30 100644 --- a/stage0/src/kernel/environment.cpp +++ b/stage0/src/kernel/environment.cpp @@ -292,7 +292,8 @@ environment environment::add(declaration const & d, bool check) const { lean_unreachable(); } -extern "C" LEAN_EXPORT object * lean_add_decl(object * env, object * decl) { +extern "C" LEAN_EXPORT object * lean_add_decl(object * env, size_t max_heartbeat, object * decl) { + scope_max_heartbeat s(max_heartbeat); return catch_kernel_exceptions([&]() { return environment(env).add(declaration(decl, true)); }); diff --git a/stage0/src/library/compiler/lcnf.cpp b/stage0/src/library/compiler/lcnf.cpp index af2af4245272..9bcd88d4cf59 100644 --- a/stage0/src/library/compiler/lcnf.cpp +++ b/stage0/src/library/compiler/lcnf.cpp @@ -266,6 +266,23 @@ class to_lcnf_fn { return e; } + unsigned get_constructor_non_prop_nfields(name ctor, unsigned nparams) { + local_ctx lctx; + expr type = env().get(ctor).get_type(); + for (unsigned i = 0; i < nparams; i++) { + lean_assert(is_pi(type)); + expr local = lctx.mk_local_decl(ngen(), binding_name(type), binding_domain(type), binding_info(type)); + type = instantiate(binding_body(type), local); + } + unsigned nfields = 0; + while (is_pi(type)) { + if (!type_checker(m_st, lctx).is_prop(binding_domain(type))) nfields++; + expr local = lctx.mk_local_decl(ngen(), binding_name(type), binding_domain(type), binding_info(type)); + type = instantiate(binding_body(type), local); + } + return nfields; + } + expr visit_no_confusion(expr const & fn, buffer & args, bool root) { name const & no_confusion_name = const_name(fn); name const & I_name = no_confusion_name.get_prefix(); @@ -297,7 +314,7 @@ class to_lcnf_fn { lean_assert(args.size() >= basic_arity + 1); unsigned major_idx = basic_arity; expr major = args[major_idx]; - unsigned nfields = get_constructor_nfields(*lhs_constructor); + unsigned nfields = get_constructor_non_prop_nfields(*lhs_constructor, nparams); while (nfields > 0) { if (!is_lambda(major)) major = eta_expand(major, nfields); diff --git a/stage0/src/runtime/interrupt.cpp b/stage0/src/runtime/interrupt.cpp index 13c242582a02..6a0751ad5717 100644 --- a/stage0/src/runtime/interrupt.cpp +++ b/stage0/src/runtime/interrupt.cpp @@ -26,7 +26,7 @@ size_t get_max_heartbeat() { return g_max_heartbeat; } void set_max_heartbeat_thousands(unsigned max) { g_max_heartbeat = static_cast(max) * 1000; } scope_heartbeat::scope_heartbeat(size_t max):flet(g_heartbeat, max) {} -scope_max_heartbeat::scope_max_heartbeat(size_t max):flet(g_max_heartbeat, max) {} +LEAN_EXPORT scope_max_heartbeat::scope_max_heartbeat(size_t max):flet(g_max_heartbeat, max) {} // separate definition to allow breakpoint in debugger void throw_heartbeat_exception() { diff --git a/stage0/src/runtime/interrupt.h b/stage0/src/runtime/interrupt.h index ccc679319305..0f6b10cc5195 100644 --- a/stage0/src/runtime/interrupt.h +++ b/stage0/src/runtime/interrupt.h @@ -35,9 +35,9 @@ LEAN_EXPORT void set_max_heartbeat_thousands(unsigned max); LEAN_EXPORT size_t get_max_heartbeat(); /* Update the thread local max heartbeat */ -class scope_max_heartbeat : flet { +class LEAN_EXPORT scope_max_heartbeat : flet { public: - scope_max_heartbeat(size_t max); + LEAN_EXPORT scope_max_heartbeat(size_t max); }; LEAN_EXPORT void check_heartbeat(); diff --git a/stage0/src/stdlib_flags.h b/stage0/src/stdlib_flags.h index 158ed2fe6013..0699845ba452 100644 --- a/stage0/src/stdlib_flags.h +++ b/stage0/src/stdlib_flags.h @@ -1,7 +1,5 @@ #include "util/options.h" -// please auto update stage0 - namespace lean { options get_default_options() { options opts; diff --git a/stage0/stdlib/Init/Control/Lawful/Basic.c b/stage0/stdlib/Init/Control/Lawful/Basic.c index 9a28244674ab..7c205ebcb17e 100644 --- a/stage0/stdlib/Init/Control/Lawful/Basic.c +++ b/stage0/stdlib/Init/Control/Lawful/Basic.c @@ -32,10 +32,12 @@ static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____clos static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__15; LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__6; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1757_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__22; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1725_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__7; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__4; +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1653_; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__16; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; @@ -45,14 +47,12 @@ static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____clos static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__31; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__28; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__25; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__13; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__18; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1693_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__32; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_; -LEAN_EXPORT lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__24; static lean_object* l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__3; static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__1() { @@ -410,7 +410,7 @@ x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1653_() { _start: { lean_object* x_1; @@ -418,7 +418,7 @@ x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1693_() { _start: { lean_object* x_1; @@ -426,7 +426,7 @@ x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1725_() { _start: { lean_object* x_1; @@ -434,7 +434,7 @@ x_1 = l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33; return x_1; } } -static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_() { +static lean_object* _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1757_() { _start: { lean_object* x_1; @@ -523,14 +523,14 @@ l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33 = _init_l___a lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631____closed__33); l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_(); lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1631_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1691_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1769_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1839_); -l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_(); -lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1909_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1653_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1653_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1653_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1693_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1693_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1693_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1725_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1725_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1725_); +l___auto____x40_Init_Control_Lawful_Basic___hyg_1757_ = _init_l___auto____x40_Init_Control_Lawful_Basic___hyg_1757_(); +lean_mark_persistent(l___auto____x40_Init_Control_Lawful_Basic___hyg_1757_); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Init/GetElem.c b/stage0/stdlib/Init/GetElem.c index 16dd879f8dbe..7872df031439 100644 --- a/stage0/stdlib/Init/GetElem.c +++ b/stage0/stdlib/Init/GetElem.c @@ -89,6 +89,7 @@ static lean_object* l___aux__Init__GetElem______macroRules__term_____x5b___x5d__ LEAN_EXPORT lean_object* l___aux__Init__GetElem______macroRules__term_____x5b___x5d_x27____1(lean_object*, lean_object*, lean_object*); lean_object* l_panic___rarg(lean_object*, lean_object*); static lean_object* l_term_____x5b___x5d___x3f___closed__12; +LEAN_EXPORT lean_object* l___auto____x40_Init_GetElem___hyg_1019_; static lean_object* l_term_____x5b___x5d___closed__1; static lean_object* l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__9; static lean_object* l_Lean_Syntax_instGetElemNatTrue___closed__2; @@ -116,7 +117,6 @@ static lean_object* l___aux__Init__GetElem______macroRules__term_____x5b___x5d__ static lean_object* l_Lean_Syntax_instGetElemNatTrue___closed__3; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__GetElem______macroRules__term_____x5b___x5d___x21__1___closed__5; -LEAN_EXPORT lean_object* l___auto____x40_Init_GetElem___hyg_1057_; static lean_object* l_term_____x5b___x5d___closed__18; static lean_object* l___auto____x40_Init_GetElem___hyg_1018____closed__21; static lean_object* l___aux__Init__GetElem______macroRules__term_____x5b___x5d__1___closed__24; @@ -1850,7 +1850,7 @@ x_1 = l___auto____x40_Init_GetElem___hyg_1018____closed__24; return x_1; } } -static lean_object* _init_l___auto____x40_Init_GetElem___hyg_1057_() { +static lean_object* _init_l___auto____x40_Init_GetElem___hyg_1019_() { _start: { lean_object* x_1; @@ -2808,8 +2808,8 @@ l___auto____x40_Init_GetElem___hyg_1018____closed__24 = _init_l___auto____x40_In lean_mark_persistent(l___auto____x40_Init_GetElem___hyg_1018____closed__24); l___auto____x40_Init_GetElem___hyg_1018_ = _init_l___auto____x40_Init_GetElem___hyg_1018_(); lean_mark_persistent(l___auto____x40_Init_GetElem___hyg_1018_); -l___auto____x40_Init_GetElem___hyg_1057_ = _init_l___auto____x40_Init_GetElem___hyg_1057_(); -lean_mark_persistent(l___auto____x40_Init_GetElem___hyg_1057_); +l___auto____x40_Init_GetElem___hyg_1019_ = _init_l___auto____x40_Init_GetElem___hyg_1019_(); +lean_mark_persistent(l___auto____x40_Init_GetElem___hyg_1019_); l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__1 = _init_l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__1(); lean_mark_persistent(l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__1); l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__2 = _init_l_Fin___aux__Init__GetElem______macroRules__tacticGet__elem__tactic__trivial__1___closed__2(); diff --git a/stage0/stdlib/Init/System/IO.c b/stage0/stdlib/Init/System/IO.c index 9e2b0066a607..7c3651f3c94a 100644 --- a/stage0/stdlib/Init/System/IO.c +++ b/stage0/stdlib/Init/System/IO.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20; LEAN_EXPORT lean_object* l_IO_FS_Handle_lock___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instInhabitedStream___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_unsafeIO___rarg(lean_object*); @@ -26,10 +25,12 @@ lean_object* lean_byte_array_copy_slice(lean_object*, lean_object*, lean_object* static lean_object* l_termPrintln_x21_______closed__9; LEAN_EXPORT lean_object* l_EIO_toIO_x27(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_eprint___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514_(lean_object*, lean_object*); lean_object* lean_io_cancel(lean_object*, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__20; static lean_object* l_IO_FS_instInhabitedSystemTime___closed__1; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7; +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202_(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422_(uint8_t, uint8_t); lean_object* lean_io_prim_handle_lock(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_getNumHeartbeats___boxed(lean_object*); LEAN_EXPORT lean_object* l_EIO_toBaseIO___rarg(lean_object*, lean_object*); @@ -39,6 +40,7 @@ LEAN_EXPORT lean_object* l_termPrintln_x21____; static lean_object* l_IO_FS_instInhabitedStream___closed__4; static lean_object* l_IO_FS_instInhabitedStream___closed__2; LEAN_EXPORT lean_object* l_IO_FS_Handle_readBinToEnd(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3; lean_object* lean_io_prim_handle_unlock(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_EIO_toIO(lean_object*, lean_object*); @@ -47,6 +49,7 @@ lean_object* lean_io_prim_handle_put_str(lean_object*, lean_object*, lean_object LEAN_EXPORT lean_object* l_IO_toEIO(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___rarg___lambda__2(lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_IO_Process_output___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2; LEAN_EXPORT lean_object* l_System_FilePath_walkDir(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_runEval(lean_object*); static lean_object* l_IO_FS_Stream_ofBuffer___elambda__3___closed__1; @@ -62,47 +65,49 @@ LEAN_EXPORT lean_object* l_MonadExcept_orElse___at_instOrElseEIO___spec__1___rar lean_object* lean_io_remove_file(lean_object*, lean_object*); static uint32_t l_IO_AccessRight_flags___closed__11; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__11; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7; LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__5; LEAN_EXPORT lean_object* l_IO_println(lean_object*); lean_object* lean_io_prim_handle_read(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_readFile___boxed(lean_object*, lean_object*); lean_object* lean_io_check_canceled(lean_object*); LEAN_EXPORT lean_object* l_IO_appPath___boxed(lean_object*); LEAN_EXPORT lean_object* l_EIO_toIO___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4; lean_object* lean_io_has_finished(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3; static uint32_t l_IO_FS_instInhabitedStream___lambda__1___closed__1; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3; LEAN_EXPORT lean_object* l_EIO_asTask(lean_object*, lean_object*); lean_object* l_EStateM_instInhabited___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_toIO_x27___rarg(lean_object*, lean_object*); lean_object* lean_get_set_stdout(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6; static lean_object* l_termPrintln_x21_______closed__8; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__3; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7; -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552_(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1; LEAN_EXPORT lean_object* l_Lean_instEvalUnit(lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__17; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__6(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_mkRef(lean_object*); lean_object* lean_io_as_task(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_runEval___rarg___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_String_quote(lean_object*); lean_object* lean_io_prim_handle_flush(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__27; LEAN_EXPORT lean_object* l_EIO_bindTask___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_getLine___boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21; static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__14; LEAN_EXPORT lean_object* l_IO_Process_run___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FileRight_other___default; LEAN_EXPORT lean_object* l_IO_FS_rename___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20; LEAN_EXPORT lean_object* l_IO_FileRight_flags___boxed(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15; LEAN_EXPORT lean_object* l_IO_hasFinished___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_mapTasks___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); @@ -121,27 +126,22 @@ uint8_t l_ByteArray_isEmpty(lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__13; lean_object* lean_io_process_child_take_stdin(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_appDir___closed__1; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16; static lean_object* l_termPrintln_x21_______closed__1; lean_object* lean_io_rename(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Mode_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10; lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__5___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_instMonadLiftSTRealWorldBaseIO(lean_object*); lean_object* lean_io_metadata(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_wait___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_runEval___spec__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13; LEAN_EXPORT lean_object* l_IO_iterate(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460____boxed(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5; LEAN_EXPORT lean_object* l_IO_FS_Mode_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_bindTask___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_System_FilePath_metadata___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____boxed(lean_object*, lean_object*); lean_object* lean_io_getenv(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__1(lean_object*, lean_object*); lean_object* l_ByteArray_extract(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__4; @@ -149,6 +149,7 @@ static lean_object* l_IO_FileRight_user___default___closed__1; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__3(lean_object*, lean_object*); static uint32_t l_IO_AccessRight_flags___closed__4; LEAN_EXPORT lean_object* l_IO_withStderr___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275_(uint8_t, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__16; lean_object* lean_io_remove_dir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_DirEntry_path(lean_object*); @@ -160,16 +161,15 @@ static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNone static lean_object* l_termPrintln_x21_______closed__13; static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__9; LEAN_EXPORT lean_object* l_allocprof___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instInhabitedSystemTime; LEAN_EXPORT lean_object* l_Lean_instEval(lean_object*); static lean_object* l_Lean_instEvalUnit___rarg___closed__2; static lean_object* l_termPrintln_x21_______closed__3; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__10; LEAN_EXPORT lean_object* l_IO_FS_FileType_toCtorIdx(uint8_t); 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_ByteArray_findIdx_x3f_loop___at_IO_FS_Stream_ofBuffer___spec__1___boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7; lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instReprSystemTime; @@ -182,44 +182,40 @@ static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202__ static uint32_t l_IO_AccessRight_flags___closed__12; LEAN_EXPORT lean_object* l_Lean_instEvalUnit___rarg(uint8_t, lean_object*); static uint32_t l_IO_AccessRight_flags___closed__8; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17; +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____boxed(lean_object*, lean_object*); lean_object* l_EStateM_instMonad(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4; static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__3; static lean_object* l_IO_FS_Stream_ofBuffer___elambda__3___closed__2; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26; LEAN_EXPORT lean_object* l_IO_getStdout___boxed(lean_object*); static lean_object* l_IO_FS_instInhabitedStream___closed__3; LEAN_EXPORT lean_object* l_System_FilePath_walkDir_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_FileType_noConfusion(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1; LEAN_EXPORT lean_object* l_IO_FS_Handle_readBinToEnd___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Mode_noConfusion___rarg___lambda__1___boxed(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11; lean_object* l_ByteArray_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ByteArray_findIdx_x3f_loop___at_IO_FS_Stream_ofBuffer___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_readBinToEnd_loop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_lines(lean_object*, lean_object*); lean_object* lean_dbg_sleep(uint32_t, lean_object*); LEAN_EXPORT lean_object* l_System_FilePath_pathExists(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__29; -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_MonadExcept_orElse___at_instOrElseEIO___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_removeDir___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_Buffer_pos___default; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__32; static lean_object* l_IO_FS_instReprFileType___closed__1; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__26; lean_object* lean_string_push(lean_object*, uint32_t); LEAN_EXPORT lean_object* l_IO_FS_instInhabitedStream; uint8_t lean_string_validate_utf8(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12; LEAN_EXPORT lean_object* l_instMonadExceptOfEIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_instEval___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instEvalUnit___boxed(lean_object*); LEAN_EXPORT lean_object* l_IO_iterate___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__2(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_cancel___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__5; LEAN_EXPORT lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202_; @@ -233,10 +229,11 @@ static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21__ LEAN_EXPORT uint32_t l_IO_AccessRight_flags(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_writeBinFile(lean_object*, lean_object*, lean_object*); lean_object* lean_io_process_child_wait(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23; LEAN_EXPORT lean_object* l_IO_println___at_Lean_instEval___spec__1(lean_object*, lean_object*); static lean_object* l_IO_Process_output___closed__1; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__20; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_EIO_mapTask(lean_object*, lean_object*, lean_object*); uint8_t l_String_isEmpty(lean_object*); @@ -264,41 +261,39 @@ lean_object* l_String_dropRight(lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__25; LEAN_EXPORT lean_object* l_IO_FS_Handle_truncate___boxed(lean_object*, lean_object*); static lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__1___closed__2; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19; LEAN_EXPORT lean_object* l_System_FilePath_isDir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_Stdio_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__4___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4; LEAN_EXPORT lean_object* l_IO_waitAny___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6; LEAN_EXPORT lean_object* l_Lean_instEvalOfRepr(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_eprintln___at___private_Init_System_IO_0__IO_eprintlnAux___spec__1(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* lean_io_timeit(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_putStrLn(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23; extern lean_object* l_ByteArray_empty; LEAN_EXPORT lean_object* l_System_FilePath_walkDir_go(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24; LEAN_EXPORT lean_object* l_IO_mapTasks___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_instBEqSystemTime___closed__1; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__4___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12; LEAN_EXPORT lean_object* l_EIO_mapTask___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8; LEAN_EXPORT lean_object* l_EIO_mapTasks___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_run___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25; static lean_object* l_IO_Process_run___closed__1; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1; LEAN_EXPORT lean_object* l_instMonadEIO(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_System_FilePath_walkDir_go___spec__1___closed__1; static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__12; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__1; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3; lean_object* lean_io_wait_any(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_tryLock___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__9; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__4; LEAN_EXPORT lean_object* l_IO_FS_withFile(lean_object*); @@ -308,17 +303,17 @@ static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__5; LEAN_EXPORT lean_object* l_BaseIO_mapTasks___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_lazyPure___rarg(lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__10; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20; LEAN_EXPORT lean_object* l_IO_sleep(uint32_t, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_readToEnd_loop(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_removeDirAll___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_getCurrentDir___boxed(lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240_(lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_timeit___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_AccessRight_flags___boxed(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__5(lean_object*, size_t, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__8; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_mk___boxed(lean_object*, lean_object*, lean_object*); @@ -326,11 +321,13 @@ LEAN_EXPORT lean_object* l_IO_FS_Handle_isTty___boxed(lean_object*, lean_object* LEAN_EXPORT lean_object* l_IO_FS_Handle_readToEnd_loop___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_print(lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__28; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5; static lean_object* l_instMonadFinallyEIO___closed__1; LEAN_EXPORT lean_object* l_IO_Process_Stdio_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_appDir(lean_object*); LEAN_EXPORT lean_object* l_IO_withStderr___at_Lean_runEval___spec__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__11; lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -355,12 +352,13 @@ static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNone lean_object* lean_get_stderr(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instInhabitedStream___lambda__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_createDirAll___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2; LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5; LEAN_EXPORT lean_object* l_IO_Process_output(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_exit___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_print___at_IO_println___spec__1(lean_object*, lean_object*); lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6; LEAN_EXPORT lean_object* l___auto____x40_Init_System_IO___hyg_1378_; LEAN_EXPORT lean_object* l_IO_Process_SpawnArgs_env___default; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -370,22 +368,22 @@ LEAN_EXPORT lean_object* l_IO_FS_Handle_flush___boxed(lean_object*, lean_object* LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer(lean_object*); LEAN_EXPORT lean_object* l_IO_print___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_withIsolatedStreams___rarg___closed__1; -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495_(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdout(lean_object*, lean_object*); lean_object* lean_io_prim_handle_write(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_Child_wait___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadBaseIO; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_prim_handle_mk(lean_object*, uint8_t, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__1; LEAN_EXPORT lean_object* l_EIO_bindTask___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14; static lean_object* l_IO_FS_instReprMetadata___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_System_FilePath_walkDir_go___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instEvalIO___rarg(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___rarg___lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14; lean_object* lean_get_set_stderr(lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__3; LEAN_EXPORT lean_object* l_EIO_toBaseIO(lean_object*, lean_object*); @@ -397,10 +395,10 @@ static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21__ extern lean_object* l_Std_Format_defWidth; static lean_object* l_Lean_instEvalUnit___rarg___closed__1; lean_object* lean_string_from_utf8(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__5(lean_object*, size_t, lean_object*); static lean_object* l_instMonadEIO___closed__1; LEAN_EXPORT uint8_t l_IO_Process_StdioConfig_stdin___default; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__18; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_runEval___spec__1(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_toEIO___rarg(lean_object*, lean_object*, lean_object*); @@ -408,7 +406,6 @@ LEAN_EXPORT lean_object* l_IO_setStderr___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_getRandomBytes___boxed(lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__2; lean_object* lean_usize_to_nat(size_t); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15; LEAN_EXPORT lean_object* l_IO_FS_instReprDirEntry; LEAN_EXPORT lean_object* l_IO_FS_Mode_noConfusion___rarg___lambda__1(lean_object*); static lean_object* l_IO_FS_instBEqFileType___closed__1; @@ -416,8 +413,8 @@ lean_object* lean_io_process_set_current_dir(lean_object*, lean_object*); static lean_object* l_IO_FS_Mode_noConfusion___rarg___closed__1; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__17; lean_object* lean_io_prim_handle_get_line(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15; LEAN_EXPORT uint8_t l_IO_AccessRight_write___default; -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746_(lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__4; static lean_object* l_IO_FS_Stream_ofBuffer___closed__1; static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__12; @@ -431,6 +428,7 @@ LEAN_EXPORT lean_object* l_Lean_runEval___rarg(lean_object*, lean_object*, lean_ LEAN_EXPORT uint8_t l_IO_Process_StdioConfig_stderr___default; LEAN_EXPORT lean_object* l_unsafeBaseIO___rarg(lean_object*); static lean_object* l_termPrintln_x21_______closed__15; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_IO_FS_removeDirAll___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_createDir___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_mapTask(lean_object*, lean_object*); @@ -438,11 +436,8 @@ LEAN_EXPORT lean_object* l_EIO_mapTasks(lean_object*, lean_object*, lean_object* LEAN_EXPORT lean_object* l_BaseIO_bindTask___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___lambda__1(lean_object*, lean_object*); static lean_object* l_termPrintln_x21_______closed__5; -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_IO_FS_writeFile___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_IO_Process_SpawnArgs_setsid___default; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__12; LEAN_EXPORT lean_object* l_System_FilePath_walkDir_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static uint32_t l_IO_AccessRight_flags___closed__1; @@ -452,10 +447,12 @@ static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__7; LEAN_EXPORT lean_object* l_IO_withStdout___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_withFile___rarg(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18; LEAN_EXPORT lean_object* l_BaseIO_toIO(lean_object*); static lean_object* l_IO_Process_run___closed__2; LEAN_EXPORT lean_object* l_IO_FS_lines___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__4; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__24; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_IO_FS_removeDirAll___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -463,6 +460,8 @@ LEAN_EXPORT lean_object* l_IO_getEnv___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdout___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__19; lean_object* lean_string_length(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16; static lean_object* l_termPrintln_x21_______closed__16; LEAN_EXPORT lean_object* l_Lean_instEvalIO(lean_object*); LEAN_EXPORT lean_object* l_EIO_bindTask(lean_object*, lean_object*, lean_object*); @@ -471,7 +470,7 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__1___boxed(lean_objec LEAN_EXPORT lean_object* l_IO_Process_SpawnArgs_args___default; static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__5; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__6; static lean_object* l_IO_FS_withIsolatedStreams___rarg___closed__2; @@ -479,15 +478,16 @@ LEAN_EXPORT lean_object* l_IO_FS_instInhabitedStream___lambda__2___boxed(lean_ob LEAN_EXPORT lean_object* l_IO_FS_createDirAll___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__13; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; uint8_t lean_uint32_dec_eq(uint32_t, uint32_t); extern lean_object* l_Task_Priority_dedicated; static lean_object* l_IO_withStdin___rarg___lambda__3___closed__1; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__6; LEAN_EXPORT lean_object* l_IO_setAccessRights___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_putStr___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_bindTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3; LEAN_EXPORT lean_object* l_EIO_bindTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_Process_setCurrentDir___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_mapTasks___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -497,23 +497,21 @@ static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21__ LEAN_EXPORT lean_object* l_IO_FS_instReprFileType; lean_object* lean_io_app_path(lean_object*); uint32_t l_String_back(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9; LEAN_EXPORT lean_object* l_System_FilePath_isDir___boxed(lean_object*, lean_object*); static lean_object* l_instMonadExceptOfEIO___closed__1; LEAN_EXPORT lean_object* l_IO_Process_getPID___boxed(lean_object*); static uint32_t l_IO_AccessRight_flags___closed__10; LEAN_EXPORT lean_object* l_IO_mapTask___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22; LEAN_EXPORT lean_object* l_BaseIO_asTask___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__23; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__2; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__12; lean_object* lean_io_realpath(lean_object*, lean_object*); lean_object* lean_string_to_utf8(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18; -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____boxed(lean_object*, lean_object*); lean_object* lean_io_exit(uint8_t, lean_object*); static lean_object* l_termPrintln_x21_______closed__10; LEAN_EXPORT lean_object* l_IO_Prim_setAccessRights___boxed(lean_object*, lean_object*, lean_object*); @@ -522,12 +520,12 @@ static lean_object* l_termPrintln_x21_______closed__2; lean_object* l_Repr_addAppParen(lean_object*, lean_object*); uint32_t lean_uint32_lor(uint32_t, uint32_t); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__16; +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____boxed(lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__22; uint32_t lean_uint32_shift_left(uint32_t, uint32_t); static lean_object* l_IO_FS_instInhabitedStream___lambda__1___closed__2; lean_object* lean_chmod(lean_object*, uint32_t, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__7; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3; uint8_t lean_int_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instBEqFileType; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -539,9 +537,7 @@ uint32_t lean_uint32_of_nat(lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__9; LEAN_EXPORT lean_object* l_IO_Process_Stdio_toCtorIdx___boxed(lean_object*); lean_object* lean_io_bind_task(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_System_FilePath_walkDir_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4; lean_object* lean_get_set_stdin(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_lazyPure(lean_object*); LEAN_EXPORT uint32_t l_IO_FileRight_flags(lean_object*); @@ -549,39 +545,37 @@ static lean_object* l_termPrintln_x21_______closed__14; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_Buffer_data___default; LEAN_EXPORT lean_object* l_IO_withStdin___rarg___lambda__2___boxed(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2; LEAN_EXPORT lean_object* l_IO_Process_Child_takeStdin___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5; LEAN_EXPORT lean_object* l_IO_setStdout___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_BaseIO_toIO___rarg(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2; static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__14; LEAN_EXPORT lean_object* l_IO_FileRight_user___default; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7; LEAN_EXPORT lean_object* l_instMonadLiftBaseIOEIO___rarg(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9; lean_object* l_EStateM_instMonadFinally(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instLESystemTime; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26; LEAN_EXPORT lean_object* l_BaseIO_mapTasks_go___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_initializing(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); static uint32_t l_IO_AccessRight_flags___closed__3; LEAN_EXPORT lean_object* l_Lean_instEval___rarg(lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_termPrintln_x21_______closed__17; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2; lean_object* lean_io_read_dir(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_monoNanosNow___boxed(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_unlock___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_mapTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_readBinToEnd_loop___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552____boxed(lean_object*, lean_object*); uint8_t lean_byte_array_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_unsafeEIO___rarg(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18; +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708_(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6; size_t lean_usize_add(size_t, size_t); lean_object* lean_io_current_dir(lean_object*); LEAN_EXPORT lean_object* l_IO_getStderr___boxed(lean_object*); uint8_t lean_uint32_dec_lt(uint32_t, uint32_t); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8; static lean_object* l_IO_FS_instReprDirEntry___closed__1; static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__11; lean_object* lean_array_uget(lean_object*, size_t); @@ -590,12 +584,16 @@ LEAN_EXPORT lean_object* l_IO_FS_FileType_noConfusion___rarg(uint8_t, uint8_t, l LEAN_EXPORT lean_object* l_IO_FS_instReprMetadata; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__3(lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629____boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_mapTask___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2; static lean_object* l_IO_FS_instOrdSystemTime___closed__1; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__4; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__15; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13; LEAN_EXPORT lean_object* l_IO_FS_Handle_readToEnd(lean_object*, lean_object*); static lean_object* l_IO_appDir___closed__2; uint8_t lean_int_dec_eq(lean_object*, lean_object*); @@ -614,22 +612,23 @@ static lean_object* l_termPrintln_x21_______closed__11; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__16; lean_object* lean_io_prim_handle_truncate(lean_object*, lean_object*); static lean_object* l_instMonadExceptOfEIO___closed__2; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6; LEAN_EXPORT lean_object* l_IO_setStdin___boxed(lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__8; LEAN_EXPORT lean_object* l_IO_FS_Mode_toCtorIdx(uint8_t); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4; LEAN_EXPORT lean_object* l_IO_asTask___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_eprintln(lean_object*); lean_object* lean_io_wait(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9; static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__6; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__1; -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313_(uint8_t, lean_object*); lean_object* lean_array_get_size(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19; LEAN_EXPORT lean_object* l_IO_FS_writeBinFile___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10; LEAN_EXPORT lean_object* l_IO_eprint___at_IO_eprintln___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9; LEAN_EXPORT lean_object* l_IO_asTask(lean_object*); LEAN_EXPORT lean_object* l_BaseIO_toEIO___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_System_FilePath_pathExists___boxed(lean_object*, lean_object*); @@ -639,6 +638,9 @@ LEAN_EXPORT lean_object* l_IO_withStdin___rarg(lean_object*, lean_object*, lean_ static lean_object* l_IO_FS_instReprSystemTime___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); static uint32_t l_IO_AccessRight_flags___closed__7; +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591____boxed(lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8; LEAN_EXPORT lean_object* l_Lean_instEvalIO___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__17; LEAN_EXPORT lean_object* l_IO_sleep___lambda__1___boxed(lean_object*, lean_object*); @@ -646,22 +648,20 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadLiftBaseIOEIO(lean_object*, lean_object*); lean_object* lean_io_get_random_bytes(size_t, lean_object*); lean_object* lean_io_process_get_pid(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5; LEAN_EXPORT lean_object* l_IO_FS_readBinFile(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instEvalOfRepr___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_EIO_asTask___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_termPrintln_x21_______closed__6; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16; LEAN_EXPORT lean_object* l_IO_FS_Handle_rewind___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_bindTask(lean_object*, lean_object*); static lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1___closed__15; LEAN_EXPORT lean_object* l_IO_FS_Stream_ofHandle___elambda__2___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1; LEAN_EXPORT lean_object* l_Lean_instEvalBaseIO(lean_object*); LEAN_EXPORT lean_object* l_IO_withStdin___at_Lean_runEval___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_instMonadLiftSTRealWorldBaseIO___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_IO_Process_StdioConfig_stdout___default; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__6; static lean_object* l_IO_tacticNonempty__list____x40_Init_System_IO___hyg_1202____closed__3; LEAN_EXPORT lean_object* l_IO___aux__Init__System__IO______macroRules__IO__tacticNonempty__list__1(lean_object*, lean_object*, lean_object*); extern uint8_t l_System_Platform_isWindows; @@ -674,20 +674,20 @@ lean_object* l_EStateM_nonBacktrackable(lean_object*); lean_object* l_String_toSubstring_x27(lean_object*); LEAN_EXPORT lean_object* l_IO_Process_run(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_setAccessRights(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams(lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__9; lean_object* lean_io_process_child_kill(lean_object*, lean_object*, lean_object*); static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__10; LEAN_EXPORT lean_object* l_IO_Process_Child_kill___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_io_prim_handle_rewind(lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_ofBuffer___closed__2; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1; +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22; lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__18; LEAN_EXPORT lean_object* l_EIO_mapTask___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_instMonadFinallyEIO(lean_object*); +static lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1; static lean_object* l___aux__Init__System__IO______macroRules__termPrintln_x21______1___closed__14; static lean_object* l_termPrintln_x21_______closed__7; static uint32_t l_IO_AccessRight_flags___closed__9; @@ -698,12 +698,12 @@ static lean_object* l___auto____x40_Init_System_IO___hyg_1378____closed__13; LEAN_EXPORT lean_object* l_IO_FS_withFile___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_ofBuffer___elambda__1(lean_object*); lean_object* lean_io_create_dir(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591_(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_putStrLn___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4; LEAN_EXPORT lean_object* l_EIO_catchExceptions(lean_object*, lean_object*); lean_object* lean_io_process_get_current_dir(lean_object*); LEAN_EXPORT lean_object* l_IO_FS_instLTSystemTime; -static lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5; LEAN_EXPORT lean_object* l_Lean_instEvalOfRepr___rarg(lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Handle_read___boxed(lean_object*, lean_object*, lean_object*); static lean_object* _init_l_instMonadEIO___closed__1() { @@ -3938,7 +3938,7 @@ x_7 = lean_apply_2(x_4, x_6, x_3); return x_7; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1() { _start: { lean_object* x_1; @@ -3946,29 +3946,29 @@ x_1 = lean_mk_string_from_bytes("root", 4); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2; +x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__4() { _start: { lean_object* x_1; @@ -3976,29 +3976,29 @@ x_1 = lean_mk_string_from_bytes(" := ", 4); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____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; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7() { _start: { lean_object* x_1; lean_object* x_2; @@ -4007,7 +4007,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8() { _start: { lean_object* x_1; @@ -4015,17 +4015,17 @@ x_1 = lean_mk_string_from_bytes("FilePath.mk ", 12); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10() { _start: { lean_object* x_1; @@ -4033,17 +4033,17 @@ x_1 = lean_mk_string_from_bytes(",", 1); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12() { _start: { lean_object* x_1; @@ -4051,17 +4051,17 @@ x_1 = lean_mk_string_from_bytes("fileName", 8); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14() { _start: { lean_object* x_1; lean_object* x_2; @@ -4070,7 +4070,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15() { _start: { lean_object* x_1; @@ -4078,35 +4078,35 @@ x_1 = lean_mk_string_from_bytes("{ ", 2); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15; x_2 = lean_string_length(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16; x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19() { _start: { lean_object* x_1; @@ -4114,17 +4114,17 @@ x_1 = lean_mk_string_from_bytes(" }", 2); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19; +x_1 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19; 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202_(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; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; @@ -4132,13 +4132,13 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = l_String_quote(x_3); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9; +x_6 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9; x_7 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); x_8 = lean_unsigned_to_nat(0u); x_9 = l_Repr_addAppParen(x_7, x_8); -x_10 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7; +x_10 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7; x_11 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); @@ -4146,11 +4146,11 @@ 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6; +x_14 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6; 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11; +x_16 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); @@ -4158,11 +4158,11 @@ x_18 = lean_box(1); 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13; +x_20 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13; 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 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_22 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5; x_23 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_23, 0, x_21); lean_ctor_set(x_23, 1, x_22); @@ -4170,7 +4170,7 @@ x_24 = lean_ctor_get(x_1, 1); x_25 = l_String_quote(x_24); x_26 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_26, 0, x_25); -x_27 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14; +x_27 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14; x_28 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_28, 0, x_27); lean_ctor_set(x_28, 1, x_26); @@ -4180,15 +4180,15 @@ lean_ctor_set_uint8(x_29, sizeof(void*)*1, x_12); x_30 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_30, 0, x_23); lean_ctor_set(x_30, 1, x_29); -x_31 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18; +x_31 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18; x_32 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_32, 0, x_31); lean_ctor_set(x_32, 1, x_30); -x_33 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20; +x_33 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20; 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17; +x_35 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17; x_36 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_36, 0, x_35); lean_ctor_set(x_36, 1, x_34); @@ -4198,11 +4198,11 @@ lean_ctor_set_uint8(x_37, sizeof(void*)*1, x_12); return x_37; } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240_(x_1, x_2); +x_3 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -4212,7 +4212,7 @@ static lean_object* _init_l_IO_FS_instReprDirEntry___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____boxed), 2, 0); return x_1; } } @@ -4306,7 +4306,7 @@ x_6 = l_IO_FS_FileType_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1() { _start: { lean_object* x_1; @@ -4314,17 +4314,17 @@ x_1 = lean_mk_string_from_bytes("IO.FS.FileType.dir", 18); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -4333,23 +4333,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4357,7 +4357,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6() { _start: { lean_object* x_1; lean_object* x_2; @@ -4366,23 +4366,23 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4390,7 +4390,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9() { _start: { lean_object* x_1; @@ -4398,33 +4398,33 @@ x_1 = lean_mk_string_from_bytes("IO.FS.FileType.file", 19); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4432,23 +4432,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4456,7 +4456,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15() { _start: { lean_object* x_1; @@ -4464,33 +4464,33 @@ x_1 = lean_mk_string_from_bytes("IO.FS.FileType.symlink", 22); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4498,23 +4498,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4522,7 +4522,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21() { _start: { lean_object* x_1; @@ -4530,33 +4530,33 @@ x_1 = lean_mk_string_from_bytes("IO.FS.FileType.other", 20); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4564,23 +4564,23 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6; -x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6; +x_2 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22; x_3 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26() { _start: { lean_object* x_1; uint8_t x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25; +x_1 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25; x_2 = 0; x_3 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_3, 0, x_1); @@ -4588,7 +4588,7 @@ lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313_(uint8_t x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275_(uint8_t x_1, lean_object* x_2) { _start: { switch (x_1) { @@ -4600,14 +4600,14 @@ x_4 = lean_nat_dec_le(x_3, x_2); if (x_4 == 0) { lean_object* x_5; lean_object* x_6; -x_5 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5; +x_5 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5; x_6 = l_Repr_addAppParen(x_5, x_2); return x_6; } else { lean_object* x_7; lean_object* x_8; -x_7 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8; +x_7 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8; x_8 = l_Repr_addAppParen(x_7, x_2); return x_8; } @@ -4620,14 +4620,14 @@ x_10 = lean_nat_dec_le(x_9, x_2); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; -x_11 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12; +x_11 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12; x_12 = l_Repr_addAppParen(x_11, x_2); return x_12; } else { lean_object* x_13; lean_object* x_14; -x_13 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14; +x_13 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14; x_14 = l_Repr_addAppParen(x_13, x_2); return x_14; } @@ -4640,14 +4640,14 @@ x_16 = lean_nat_dec_le(x_15, x_2); if (x_16 == 0) { lean_object* x_17; lean_object* x_18; -x_17 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18; +x_17 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18; x_18 = l_Repr_addAppParen(x_17, x_2); return x_18; } else { lean_object* x_19; lean_object* x_20; -x_19 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20; +x_19 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20; x_20 = l_Repr_addAppParen(x_19, x_2); return x_20; } @@ -4660,14 +4660,14 @@ x_22 = lean_nat_dec_le(x_21, x_2); if (x_22 == 0) { lean_object* x_23; lean_object* x_24; -x_23 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24; +x_23 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24; x_24 = l_Repr_addAppParen(x_23, x_2); return x_24; } else { lean_object* x_25; lean_object* x_26; -x_25 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26; +x_25 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26; x_26 = l_Repr_addAppParen(x_25, x_2); return x_26; } @@ -4675,13 +4675,13 @@ return x_26; } } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; x_3 = lean_unbox(x_1); lean_dec(x_1); -x_4 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313_(x_3, x_2); +x_4 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275_(x_3, x_2); lean_dec(x_2); return x_4; } @@ -4690,7 +4690,7 @@ static lean_object* _init_l_IO_FS_instReprFileType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____boxed), 2, 0); return x_1; } } @@ -4702,7 +4702,7 @@ x_1 = l_IO_FS_instReprFileType___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -4714,7 +4714,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422____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; @@ -4722,7 +4722,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460_(x_3, x_4); +x_5 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -4731,7 +4731,7 @@ static lean_object* _init_l_IO_FS_instBEqFileType___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422____boxed), 2, 0); return x_1; } } @@ -4743,7 +4743,7 @@ x_1 = l_IO_FS_instBEqFileType___closed__1; return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1() { _start: { lean_object* x_1; @@ -4751,41 +4751,41 @@ x_1 = lean_mk_string_from_bytes("sec", 3); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1; +x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2; +x_2 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____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; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -4794,7 +4794,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__6() { _start: { lean_object* x_1; @@ -4802,17 +4802,17 @@ x_1 = lean_mk_string_from_bytes("nsec", 4); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6; +x_1 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____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_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint32_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; @@ -4820,7 +4820,7 @@ x_3 = lean_ctor_get(x_1, 0); x_4 = l_Int_repr(x_3); x_5 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_5, 0, x_4); -x_6 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5; +x_6 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -4828,11 +4828,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4; +x_10 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11; +x_12 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -4840,11 +4840,11 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7; +x_16 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7; x_17 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_18 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5; x_19 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_19, 0, x_17); lean_ctor_set(x_19, 1, x_18); @@ -4853,7 +4853,7 @@ x_21 = lean_uint32_to_nat(x_20); x_22 = l___private_Init_Data_Repr_0__Nat_reprFast(x_21); x_23 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_23, 0, x_22); -x_24 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7; +x_24 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7; x_25 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); @@ -4863,15 +4863,15 @@ lean_ctor_set_uint8(x_26, sizeof(void*)*1, x_8); x_27 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_27, 0, x_19); lean_ctor_set(x_27, 1, x_26); -x_28 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18; +x_28 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_28); lean_ctor_set(x_29, 1, x_27); -x_30 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20; +x_30 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20; x_31 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); -x_32 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17; +x_32 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17; x_33 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_33, 0, x_32); lean_ctor_set(x_33, 1, x_31); @@ -4881,11 +4881,11 @@ lean_ctor_set_uint8(x_34, sizeof(void*)*1, x_8); return x_34; } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495_(x_1, x_2); +x_3 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -4895,7 +4895,7 @@ static lean_object* _init_l_IO_FS_instReprSystemTime___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____boxed), 2, 0); return x_1; } } @@ -4907,7 +4907,7 @@ x_1 = l_IO_FS_instReprSystemTime___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint32_t x_4; lean_object* x_5; uint32_t x_6; uint8_t x_7; @@ -4930,11 +4930,11 @@ return x_9; } } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552_(x_1, x_2); +x_3 = l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -4945,7 +4945,7 @@ static lean_object* _init_l_IO_FS_instBEqSystemTime___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2552____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_beqSystemTime____x40_Init_System_IO___hyg_2514____boxed), 2, 0); return x_1; } } @@ -4957,7 +4957,7 @@ x_1 = l_IO_FS_instBEqSystemTime___closed__1; return x_1; } } -LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; uint32_t x_4; lean_object* x_5; uint32_t x_6; uint8_t x_7; @@ -5012,11 +5012,11 @@ return x_15; } } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591____boxed(lean_object* x_1, lean_object* x_2) { _start: { uint8_t x_3; lean_object* x_4; -x_3 = l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629_(x_1, x_2); +x_3 = l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); x_4 = lean_box(x_3); @@ -5027,7 +5027,7 @@ static lean_object* _init_l_IO_FS_instOrdSystemTime___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2629____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_ordSystemTime____x40_Init_System_IO___hyg_2591____boxed), 2, 0); return x_1; } } @@ -5084,7 +5084,7 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1() { _start: { lean_object* x_1; @@ -5092,41 +5092,41 @@ x_1 = lean_mk_string_from_bytes("accessed", 8); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1; +x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2; +x_2 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2; x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3; -x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3; +x_2 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____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; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__5() { _start: { lean_object* x_1; @@ -5134,17 +5134,17 @@ x_1 = lean_mk_string_from_bytes("modified", 8); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__6() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5; +x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____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___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7() { _start: { lean_object* x_1; @@ -5152,17 +5152,17 @@ x_1 = lean_mk_string_from_bytes("byteSize", 8); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7; +x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9() { _start: { lean_object* x_1; @@ -5170,24 +5170,24 @@ x_1 = lean_mk_string_from_bytes("type", 4); return x_1; } } -static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10() { +static lean_object* _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9; +x_1 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9; 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_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708_(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint64_t x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; 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; lean_object* x_52; lean_object* x_53; lean_object* x_54; x_3 = lean_ctor_get(x_1, 0); x_4 = lean_unsigned_to_nat(0u); -x_5 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495_(x_3, x_4); -x_6 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14; +x_5 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457_(x_3, x_4); +x_6 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14; x_7 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_7, 0, x_6); lean_ctor_set(x_7, 1, x_5); @@ -5195,11 +5195,11 @@ x_8 = 0; x_9 = lean_alloc_ctor(6, 1, 1); lean_ctor_set(x_9, 0, x_7); lean_ctor_set_uint8(x_9, sizeof(void*)*1, x_8); -x_10 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4; +x_10 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4; x_11 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_11, 0, x_10); lean_ctor_set(x_11, 1, x_9); -x_12 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11; +x_12 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11; x_13 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_13, 0, x_11); lean_ctor_set(x_13, 1, x_12); @@ -5207,16 +5207,16 @@ x_14 = lean_box(1); x_15 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_15, 0, x_13); lean_ctor_set(x_15, 1, x_14); -x_16 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6; +x_16 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____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___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5; +x_18 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5; 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 = lean_ctor_get(x_1, 1); -x_21 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495_(x_20, x_4); +x_21 = l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457_(x_20, x_4); x_22 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_22, 0, x_6); lean_ctor_set(x_22, 1, x_21); @@ -5232,7 +5232,7 @@ lean_ctor_set(x_25, 1, x_12); x_26 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_26, 0, x_25); lean_ctor_set(x_26, 1, x_14); -x_27 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8; +x_27 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8; x_28 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_28, 0, x_26); lean_ctor_set(x_28, 1, x_27); @@ -5259,7 +5259,7 @@ lean_ctor_set(x_37, 1, x_12); x_38 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_38, 0, x_37); lean_ctor_set(x_38, 1, x_14); -x_39 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10; +x_39 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10; x_40 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); @@ -5267,8 +5267,8 @@ x_41 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_41, 0, x_40); lean_ctor_set(x_41, 1, x_18); x_42 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 8); -x_43 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313_(x_42, x_4); -x_44 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7; +x_43 = l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275_(x_42, x_4); +x_44 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7; x_45 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_45, 0, x_44); lean_ctor_set(x_45, 1, x_43); @@ -5278,15 +5278,15 @@ lean_ctor_set_uint8(x_46, sizeof(void*)*1, x_8); x_47 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_47, 0, x_41); lean_ctor_set(x_47, 1, x_46); -x_48 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18; +x_48 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18; 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20; +x_50 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20; 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_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17; +x_52 = l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17; x_53 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_53, 0, x_52); lean_ctor_set(x_53, 1, x_51); @@ -5296,11 +5296,11 @@ lean_ctor_set_uint8(x_54, sizeof(void*)*1, x_8); return x_54; } } -LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746_(x_1, x_2); +x_3 = l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708_(x_1, x_2); lean_dec(x_2); lean_dec(x_1); return x_3; @@ -5310,7 +5310,7 @@ static lean_object* _init_l_IO_FS_instReprMetadata___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____boxed), 2, 0); return x_1; } } @@ -5356,7 +5356,7 @@ x_5 = lean_ctor_get(x_3, 0); x_6 = lean_ctor_get_uint8(x_5, sizeof(void*)*2 + 8); lean_dec(x_5); x_7 = 0; -x_8 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460_(x_6, x_7); +x_8 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422_(x_6, x_7); x_9 = lean_box(x_8); lean_ctor_set(x_3, 0, x_9); return x_3; @@ -5372,7 +5372,7 @@ lean_dec(x_3); x_12 = lean_ctor_get_uint8(x_10, sizeof(void*)*2 + 8); lean_dec(x_10); x_13 = 0; -x_14 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2460_(x_12, x_13); +x_14 = l___private_Init_System_IO_0__IO_FS_beqFileType____x40_Init_System_IO___hyg_2422_(x_12, x_13); x_15 = lean_box(x_14); x_16 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_16, 0, x_15); @@ -11282,102 +11282,102 @@ l_IO_FS_instInhabitedStream___closed__5 = _init_l_IO_FS_instInhabitedStream___cl lean_mark_persistent(l_IO_FS_instInhabitedStream___closed__5); l_IO_FS_instInhabitedStream = _init_l_IO_FS_instInhabitedStream(); lean_mark_persistent(l_IO_FS_instInhabitedStream); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__1); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__2); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__3); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__4); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__5); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__6); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__7); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__8); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__9); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__10); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__11); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__12); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__13); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__14); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__15); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__16); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__17); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__18); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__19); -l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2240____closed__20); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__1); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__2); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__3); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__4(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__4); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__5); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__6); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__7); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__8); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__9); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__10); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__11); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__12); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__13); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__14); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__15); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__16); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__17); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__18); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__19); +l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20 = _init_l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprDirEntry____x40_Init_System_IO___hyg_2202____closed__20); l_IO_FS_instReprDirEntry___closed__1 = _init_l_IO_FS_instReprDirEntry___closed__1(); lean_mark_persistent(l_IO_FS_instReprDirEntry___closed__1); l_IO_FS_instReprDirEntry = _init_l_IO_FS_instReprDirEntry(); lean_mark_persistent(l_IO_FS_instReprDirEntry); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__1); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__2); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__3); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__4); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__5); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__6); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__7); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__8); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__9); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__10); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__11); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__12); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__13); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__14); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__15); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__16); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__17); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__18); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__19); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__20); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__21); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__22); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__23); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__24); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__25); -l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2313____closed__26); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__1); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__2); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__3); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__4); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__5); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__6); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__7); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__8); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__9); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__10); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__11); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__12); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__13); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__14); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__15); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__16); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__17); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__18); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__19); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__20); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__21); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__22); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__23); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__24); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__25); +l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26 = _init_l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprFileType____x40_Init_System_IO___hyg_2275____closed__26); l_IO_FS_instReprFileType___closed__1 = _init_l_IO_FS_instReprFileType___closed__1(); lean_mark_persistent(l_IO_FS_instReprFileType___closed__1); l_IO_FS_instReprFileType = _init_l_IO_FS_instReprFileType(); @@ -11386,20 +11386,20 @@ l_IO_FS_instBEqFileType___closed__1 = _init_l_IO_FS_instBEqFileType___closed__1( lean_mark_persistent(l_IO_FS_instBEqFileType___closed__1); l_IO_FS_instBEqFileType = _init_l_IO_FS_instBEqFileType(); lean_mark_persistent(l_IO_FS_instBEqFileType); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__1); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__2); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__3); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__4); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__5); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__6); -l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2495____closed__7); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__1); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__2); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__3); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__4); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__5); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__6(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__6); +l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprSystemTime____x40_Init_System_IO___hyg_2457____closed__7); l_IO_FS_instReprSystemTime___closed__1 = _init_l_IO_FS_instReprSystemTime___closed__1(); lean_mark_persistent(l_IO_FS_instReprSystemTime___closed__1); l_IO_FS_instReprSystemTime = _init_l_IO_FS_instReprSystemTime(); @@ -11422,26 +11422,26 @@ l_IO_FS_instLTSystemTime = _init_l_IO_FS_instLTSystemTime(); lean_mark_persistent(l_IO_FS_instLTSystemTime); l_IO_FS_instLESystemTime = _init_l_IO_FS_instLESystemTime(); lean_mark_persistent(l_IO_FS_instLESystemTime); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__1); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__2); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__3); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__4); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__5); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__6); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__7); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__8); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__9); -l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10(); -lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2746____closed__10); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__1); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__2); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__3); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__4); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__5 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__5(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__5); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__6 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__6(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__6); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__7); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__8); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__9); +l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10 = _init_l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10(); +lean_mark_persistent(l___private_Init_System_IO_0__IO_FS_reprMetadata____x40_Init_System_IO___hyg_2708____closed__10); l_IO_FS_instReprMetadata___closed__1 = _init_l_IO_FS_instReprMetadata___closed__1(); lean_mark_persistent(l_IO_FS_instReprMetadata___closed__1); l_IO_FS_instReprMetadata = _init_l_IO_FS_instReprMetadata(); diff --git a/stage0/stdlib/Lean.c b/stage0/stdlib/Lean.c index b193b9755ab4..28da58c4e323 100644 --- a/stage0/stdlib/Lean.c +++ b/stage0/stdlib/Lean.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean -// Imports: Init Lean.Data Lean.Compiler Lean.Environment Lean.Modifiers Lean.ProjFns Lean.Runtime Lean.ResolveName Lean.Attributes Lean.Parser Lean.ReducibilityAttrs Lean.Elab Lean.Class Lean.LocalContext Lean.MetavarContext Lean.AuxRecursor Lean.Meta Lean.Util Lean.Eval Lean.Structure Lean.PrettyPrinter Lean.CoreM Lean.ReservedNameAction Lean.InternalExceptionId Lean.Server Lean.ScopedEnvExtension Lean.DocString Lean.DeclarationRange Lean.LazyInitExtension Lean.LoadDynlib Lean.Widget Lean.Log Lean.Linter Lean.SubExpr Lean.LabelAttribute +// Imports: Init Lean.Data Lean.Compiler Lean.Environment Lean.Modifiers Lean.ProjFns Lean.Runtime Lean.ResolveName Lean.Attributes Lean.Parser Lean.ReducibilityAttrs Lean.Elab Lean.Class Lean.LocalContext Lean.MetavarContext Lean.AuxRecursor Lean.Meta Lean.Util Lean.Eval Lean.Structure Lean.PrettyPrinter Lean.CoreM Lean.ReservedNameAction Lean.InternalExceptionId Lean.Server Lean.ScopedEnvExtension Lean.DocString Lean.DeclarationRange Lean.LazyInitExtension Lean.LoadDynlib Lean.Widget Lean.Log Lean.Linter Lean.SubExpr Lean.LabelAttribute Lean.AddDecl #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -48,6 +48,7 @@ lean_object* initialize_Lean_Log(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Linter(uint8_t builtin, lean_object*); lean_object* initialize_Lean_SubExpr(uint8_t builtin, lean_object*); lean_object* initialize_Lean_LabelAttribute(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean(uint8_t builtin, lean_object* w) { lean_object * res; @@ -158,6 +159,9 @@ lean_dec_ref(res); res = initialize_Lean_LabelAttribute(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(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/AddDecl.c b/stage0/stdlib/Lean/AddDecl.c new file mode 100644 index 000000000000..50e8ece93195 --- /dev/null +++ b/stage0/stdlib/Lean/AddDecl.c @@ -0,0 +1,2581 @@ +// Lean compiler output +// Module: Lean.AddDecl +// Imports: Lean.CoreM +#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 +static lean_object* l_Lean_addDecl___lambda__1___closed__2; +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_profiler; +lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); +extern lean_object* l_Lean_trace_profiler_useHearbeats; +lean_object* lean_add_decl(lean_object*, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2; +lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__4; +double lean_float_div(double, double); +lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3; +lean_object* l_Lean_compileDecl(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_addDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); +uint8_t lean_float_decLt(double, double); +lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___closed__2; +lean_object* lean_io_get_num_heartbeats(lean_object*); +lean_object* l_Lean_stringToMessageData(lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Environment_compileDecl(lean_object*, lean_object*, lean_object*); +size_t lean_usize_of_nat(lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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_object* lean_st_ref_take(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___closed__1; +static lean_object* l_Lean_addDecl___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); +lean_object* lean_st_ref_get(lean_object*, lean_object*); +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* lean_io_mono_nanos_now(lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8(lean_object*); +static lean_object* l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +extern lean_object* l_Lean_warningAsError; +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__1; +extern lean_object* l_Lean_trace_profiler_threshold; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2; +static lean_object* l_Lean_addDecl___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); +double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1; +static lean_object* l_Lean_addDecl___lambda__3___closed__3; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__5(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile___boxed(lean_object*, lean_object*, lean_object*); +extern lean_object* l_Lean_trace_profiler; +lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1; +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_addDecl___closed__4; +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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 double l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5; +static lean_object* l_Lean_addDecl___lambda__3___closed__1; +lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); +LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2; +lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__4(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +double lean_float_sub(double, double); +LEAN_EXPORT lean_object* l_Lean_Environment_addDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; size_t x_5; lean_object* x_6; +x_4 = l_Lean_Core_getMaxHeartbeats(x_2); +x_5 = lean_usize_of_nat(x_4); +lean_dec(x_4); +x_6 = lean_add_decl(x_1, x_5, x_3); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Environment_addDecl(x_1, x_2, x_3); +lean_dec(x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Environment_addDecl(x_1, x_2, x_3); +if (lean_obj_tag(x_4) == 0) +{ +uint8_t x_5; +lean_dec(x_3); +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) +{ +return x_4; +} +else +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_7, 0, x_6); +return x_7; +} +} +else +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_4, 0); +lean_inc(x_8); +lean_dec(x_4); +x_9 = l_Lean_Environment_compileDecl(x_8, x_2, x_3); +return x_9; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Environment_addAndCompile(x_1, x_2, x_3); +lean_dec(x_2); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__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; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(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); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +lean_inc(x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); +x_7 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_6, x_2, x_3, x_4); +lean_dec(x_2); +return x_7; +} +} +static lean_object* _init_l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_warningAsError; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__5(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; uint8_t x_265; uint8_t x_266; +x_265 = 2; +x_266 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_265); +if (x_266 == 0) +{ +lean_object* x_267; +x_267 = lean_box(0); +x_7 = x_267; +goto block_264; +} +else +{ +lean_object* x_268; uint8_t x_269; +lean_inc(x_2); +x_268 = l_Lean_MessageData_hasSyntheticSorry(x_2); +x_269 = lean_unbox(x_268); +lean_dec(x_268); +if (x_269 == 0) +{ +lean_object* x_270; +x_270 = lean_box(0); +x_7 = x_270; +goto block_264; +} +else +{ +lean_object* x_271; lean_object* x_272; +lean_dec(x_2); +x_271 = lean_box(0); +x_272 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_6); +return x_272; +} +} +block_264: +{ +uint8_t x_8; lean_object* x_258; uint8_t x_259; uint8_t x_260; +lean_dec(x_7); +x_258 = lean_ctor_get(x_4, 2); +x_259 = 1; +x_260 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_259); +if (x_260 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +lean_object* x_261; uint8_t x_262; +x_261 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2; +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); +if (x_262 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +uint8_t x_263; +x_263 = 2; +x_8 = x_263; +goto block_257; +} +} +block_257: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_ctor_get(x_4, 5); +x_12 = lean_ctor_get(x_4, 6); +x_13 = lean_ctor_get(x_4, 7); +x_14 = l_Lean_replaceRef(x_1, x_11); +x_15 = 0; +x_16 = l_Lean_Syntax_getPos_x3f(x_14, x_15); +x_17 = l_Lean_Syntax_getTailPos_x3f(x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_18 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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_unsigned_to_nat(0u); +x_22 = l_Lean_FileMap_toPosition(x_10, x_21); +lean_inc(x_22); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_inc(x_13); +lean_inc(x_12); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_13); +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +x_26 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_27 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 3, x_26); +lean_ctor_set(x_27, 4, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_8); +x_28 = lean_st_ref_take(x_5, x_20); +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; uint8_t x_35; +x_32 = lean_ctor_get(x_29, 5); +x_33 = l_Lean_PersistentArray_push___rarg(x_32, x_27); +lean_ctor_set(x_29, 5, x_33); +x_34 = lean_st_ref_set(x_5, x_29, x_30); +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_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_29, 0); +x_42 = lean_ctor_get(x_29, 1); +x_43 = lean_ctor_get(x_29, 2); +x_44 = lean_ctor_get(x_29, 3); +x_45 = lean_ctor_get(x_29, 4); +x_46 = lean_ctor_get(x_29, 5); +x_47 = lean_ctor_get(x_29, 6); +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_inc(x_41); +lean_dec(x_29); +x_48 = l_Lean_PersistentArray_push___rarg(x_46, x_27); +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_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_48); +lean_ctor_set(x_49, 6, x_47); +x_50 = lean_st_ref_set(x_5, x_49, x_30); +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 +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_17); +if (x_55 == 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; lean_object* x_69; uint8_t x_70; +x_56 = lean_ctor_get(x_17, 0); +x_57 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l_Lean_FileMap_toPosition(x_10, x_60); +x_62 = l_Lean_FileMap_toPosition(x_10, x_56); +lean_dec(x_56); +lean_ctor_set(x_17, 0, x_62); +lean_inc(x_13); +lean_inc(x_12); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_12); +lean_ctor_set(x_63, 1, x_13); +x_64 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +x_65 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_66 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_66, 0, x_9); +lean_ctor_set(x_66, 1, x_61); +lean_ctor_set(x_66, 2, x_17); +lean_ctor_set(x_66, 3, x_65); +lean_ctor_set(x_66, 4, x_64); +lean_ctor_set_uint8(x_66, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_66, sizeof(void*)*5 + 1, x_8); +x_67 = lean_st_ref_take(x_5, x_59); +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_is_exclusive(x_68); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_68, 5); +x_72 = l_Lean_PersistentArray_push___rarg(x_71, x_66); +lean_ctor_set(x_68, 5, x_72); +x_73 = lean_st_ref_set(x_5, x_68, x_69); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +x_76 = lean_box(0); +lean_ctor_set(x_73, 0, x_76); +return x_73; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_80 = lean_ctor_get(x_68, 0); +x_81 = lean_ctor_get(x_68, 1); +x_82 = lean_ctor_get(x_68, 2); +x_83 = lean_ctor_get(x_68, 3); +x_84 = lean_ctor_get(x_68, 4); +x_85 = lean_ctor_get(x_68, 5); +x_86 = lean_ctor_get(x_68, 6); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_68); +x_87 = l_Lean_PersistentArray_push___rarg(x_85, x_66); +x_88 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 1, x_81); +lean_ctor_set(x_88, 2, x_82); +lean_ctor_set(x_88, 3, x_83); +lean_ctor_set(x_88, 4, x_84); +lean_ctor_set(x_88, 5, x_87); +lean_ctor_set(x_88, 6, x_86); +x_89 = lean_st_ref_set(x_5, x_88, x_69); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_91 = x_89; +} else { + lean_dec_ref(x_89); + x_91 = lean_box(0); +} +x_92 = lean_box(0); +if (lean_is_scalar(x_91)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_91; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +return x_93; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_94 = lean_ctor_get(x_17, 0); +lean_inc(x_94); +lean_dec(x_17); +x_95 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l_Lean_FileMap_toPosition(x_10, x_98); +x_100 = l_Lean_FileMap_toPosition(x_10, x_94); +lean_dec(x_94); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +lean_inc(x_13); +lean_inc(x_12); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_12); +lean_ctor_set(x_102, 1, x_13); +x_103 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_96); +x_104 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_105 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_105, 0, x_9); +lean_ctor_set(x_105, 1, x_99); +lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 3, x_104); +lean_ctor_set(x_105, 4, x_103); +lean_ctor_set_uint8(x_105, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_105, sizeof(void*)*5 + 1, x_8); +x_106 = lean_st_ref_take(x_5, x_97); +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_ctor_get(x_107, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +x_112 = lean_ctor_get(x_107, 3); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 4); +lean_inc(x_113); +x_114 = lean_ctor_get(x_107, 5); +lean_inc(x_114); +x_115 = lean_ctor_get(x_107, 6); +lean_inc(x_115); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + lean_ctor_release(x_107, 3); + lean_ctor_release(x_107, 4); + lean_ctor_release(x_107, 5); + lean_ctor_release(x_107, 6); + x_116 = x_107; +} else { + lean_dec_ref(x_107); + x_116 = lean_box(0); +} +x_117 = l_Lean_PersistentArray_push___rarg(x_114, x_105); +if (lean_is_scalar(x_116)) { + x_118 = lean_alloc_ctor(0, 7, 0); +} else { + x_118 = x_116; +} +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_110); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set(x_118, 3, x_112); +lean_ctor_set(x_118, 4, x_113); +lean_ctor_set(x_118, 5, x_117); +lean_ctor_set(x_118, 6, x_115); +x_119 = lean_st_ref_set(x_5, x_118, x_108); +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 +{ +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_124; +x_124 = !lean_is_exclusive(x_16); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_125 = lean_ctor_get(x_16, 0); +x_126 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_FileMap_toPosition(x_10, x_125); +lean_dec(x_125); +lean_inc(x_129); +lean_ctor_set(x_16, 0, x_129); +lean_inc(x_13); +lean_inc(x_12); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_12); +lean_ctor_set(x_130, 1, x_13); +x_131 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_127); +x_132 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_133 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_133, 0, x_9); +lean_ctor_set(x_133, 1, x_129); +lean_ctor_set(x_133, 2, x_16); +lean_ctor_set(x_133, 3, x_132); +lean_ctor_set(x_133, 4, x_131); +lean_ctor_set_uint8(x_133, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_133, sizeof(void*)*5 + 1, x_8); +x_134 = lean_st_ref_take(x_5, x_128); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = !lean_is_exclusive(x_135); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_138 = lean_ctor_get(x_135, 5); +x_139 = l_Lean_PersistentArray_push___rarg(x_138, x_133); +lean_ctor_set(x_135, 5, x_139); +x_140 = lean_st_ref_set(x_5, x_135, x_136); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_140, 0); +lean_dec(x_142); +x_143 = lean_box(0); +lean_ctor_set(x_140, 0, x_143); +return x_140; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_140, 1); +lean_inc(x_144); +lean_dec(x_140); +x_145 = lean_box(0); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_147 = lean_ctor_get(x_135, 0); +x_148 = lean_ctor_get(x_135, 1); +x_149 = lean_ctor_get(x_135, 2); +x_150 = lean_ctor_get(x_135, 3); +x_151 = lean_ctor_get(x_135, 4); +x_152 = lean_ctor_get(x_135, 5); +x_153 = lean_ctor_get(x_135, 6); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_135); +x_154 = l_Lean_PersistentArray_push___rarg(x_152, x_133); +x_155 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_155, 0, x_147); +lean_ctor_set(x_155, 1, x_148); +lean_ctor_set(x_155, 2, x_149); +lean_ctor_set(x_155, 3, x_150); +lean_ctor_set(x_155, 4, x_151); +lean_ctor_set(x_155, 5, x_154); +lean_ctor_set(x_155, 6, x_153); +x_156 = lean_st_ref_set(x_5, x_155, x_136); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_158 = x_156; +} else { + lean_dec_ref(x_156); + x_158 = lean_box(0); +} +x_159 = lean_box(0); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_158; +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_161 = lean_ctor_get(x_16, 0); +lean_inc(x_161); +lean_dec(x_16); +x_162 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_FileMap_toPosition(x_10, x_161); +lean_dec(x_161); +lean_inc(x_165); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +lean_inc(x_13); +lean_inc(x_12); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_12); +lean_ctor_set(x_167, 1, x_13); +x_168 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_163); +x_169 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_170 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_170, 0, x_9); +lean_ctor_set(x_170, 1, x_165); +lean_ctor_set(x_170, 2, x_166); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_168); +lean_ctor_set_uint8(x_170, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_170, sizeof(void*)*5 + 1, x_8); +x_171 = lean_st_ref_take(x_5, x_164); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_172, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_172, 2); +lean_inc(x_176); +x_177 = lean_ctor_get(x_172, 3); +lean_inc(x_177); +x_178 = lean_ctor_get(x_172, 4); +lean_inc(x_178); +x_179 = lean_ctor_get(x_172, 5); +lean_inc(x_179); +x_180 = lean_ctor_get(x_172, 6); +lean_inc(x_180); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + lean_ctor_release(x_172, 2); + lean_ctor_release(x_172, 3); + lean_ctor_release(x_172, 4); + lean_ctor_release(x_172, 5); + lean_ctor_release(x_172, 6); + x_181 = x_172; +} else { + lean_dec_ref(x_172); + x_181 = lean_box(0); +} +x_182 = l_Lean_PersistentArray_push___rarg(x_179, x_170); +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 7, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_174); +lean_ctor_set(x_183, 1, x_175); +lean_ctor_set(x_183, 2, x_176); +lean_ctor_set(x_183, 3, x_177); +lean_ctor_set(x_183, 4, x_178); +lean_ctor_set(x_183, 5, x_182); +lean_ctor_set(x_183, 6, x_180); +x_184 = lean_st_ref_set(x_5, x_183, x_173); +x_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = lean_box(0); +if (lean_is_scalar(x_186)) { + x_188 = lean_alloc_ctor(0, 2, 0); +} else { + x_188 = x_186; +} +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_185); +return x_188; +} +} +else +{ +lean_object* x_189; uint8_t x_190; +x_189 = lean_ctor_get(x_16, 0); +lean_inc(x_189); +lean_dec(x_16); +x_190 = !lean_is_exclusive(x_17); +if (x_190 == 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; 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; uint8_t x_204; +x_191 = lean_ctor_get(x_17, 0); +x_192 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_196 = l_Lean_FileMap_toPosition(x_10, x_191); +lean_dec(x_191); +lean_ctor_set(x_17, 0, x_196); +lean_inc(x_13); +lean_inc(x_12); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_12); +lean_ctor_set(x_197, 1, x_13); +x_198 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_193); +x_199 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_200 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_200, 0, x_9); +lean_ctor_set(x_200, 1, x_195); +lean_ctor_set(x_200, 2, x_17); +lean_ctor_set(x_200, 3, x_199); +lean_ctor_set(x_200, 4, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_200, sizeof(void*)*5 + 1, x_8); +x_201 = lean_st_ref_take(x_5, x_194); +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_is_exclusive(x_202); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; +x_205 = lean_ctor_get(x_202, 5); +x_206 = l_Lean_PersistentArray_push___rarg(x_205, x_200); +lean_ctor_set(x_202, 5, x_206); +x_207 = lean_st_ref_set(x_5, x_202, x_203); +x_208 = !lean_is_exclusive(x_207); +if (x_208 == 0) +{ +lean_object* x_209; lean_object* x_210; +x_209 = lean_ctor_get(x_207, 0); +lean_dec(x_209); +x_210 = lean_box(0); +lean_ctor_set(x_207, 0, x_210); +return x_207; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +lean_dec(x_207); +x_212 = lean_box(0); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_211); +return x_213; +} +} +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; +x_214 = lean_ctor_get(x_202, 0); +x_215 = lean_ctor_get(x_202, 1); +x_216 = lean_ctor_get(x_202, 2); +x_217 = lean_ctor_get(x_202, 3); +x_218 = lean_ctor_get(x_202, 4); +x_219 = lean_ctor_get(x_202, 5); +x_220 = lean_ctor_get(x_202, 6); +lean_inc(x_220); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_202); +x_221 = l_Lean_PersistentArray_push___rarg(x_219, x_200); +x_222 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_222, 0, x_214); +lean_ctor_set(x_222, 1, x_215); +lean_ctor_set(x_222, 2, x_216); +lean_ctor_set(x_222, 3, x_217); +lean_ctor_set(x_222, 4, x_218); +lean_ctor_set(x_222, 5, x_221); +lean_ctor_set(x_222, 6, x_220); +x_223 = lean_st_ref_set(x_5, x_222, x_203); +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_223); + x_225 = lean_box(0); +} +x_226 = lean_box(0); +if (lean_is_scalar(x_225)) { + x_227 = lean_alloc_ctor(0, 2, 0); +} else { + x_227 = x_225; +} +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_224); +return x_227; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_228 = lean_ctor_get(x_17, 0); +lean_inc(x_228); +lean_dec(x_17); +x_229 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_233 = l_Lean_FileMap_toPosition(x_10, x_228); +lean_dec(x_228); +x_234 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_234, 0, x_233); +lean_inc(x_13); +lean_inc(x_12); +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_12); +lean_ctor_set(x_235, 1, x_13); +x_236 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_236, 0, x_235); +lean_ctor_set(x_236, 1, x_230); +x_237 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +lean_inc(x_9); +x_238 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_238, 0, x_9); +lean_ctor_set(x_238, 1, x_232); +lean_ctor_set(x_238, 2, x_234); +lean_ctor_set(x_238, 3, x_237); +lean_ctor_set(x_238, 4, x_236); +lean_ctor_set_uint8(x_238, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_238, sizeof(void*)*5 + 1, x_8); +x_239 = lean_st_ref_take(x_5, x_231); +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 = lean_ctor_get(x_240, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_240, 1); +lean_inc(x_243); +x_244 = lean_ctor_get(x_240, 2); +lean_inc(x_244); +x_245 = lean_ctor_get(x_240, 3); +lean_inc(x_245); +x_246 = lean_ctor_get(x_240, 4); +lean_inc(x_246); +x_247 = lean_ctor_get(x_240, 5); +lean_inc(x_247); +x_248 = lean_ctor_get(x_240, 6); +lean_inc(x_248); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + lean_ctor_release(x_240, 2); + lean_ctor_release(x_240, 3); + lean_ctor_release(x_240, 4); + lean_ctor_release(x_240, 5); + lean_ctor_release(x_240, 6); + x_249 = x_240; +} else { + lean_dec_ref(x_240); + x_249 = lean_box(0); +} +x_250 = l_Lean_PersistentArray_push___rarg(x_247, x_238); +if (lean_is_scalar(x_249)) { + x_251 = lean_alloc_ctor(0, 7, 0); +} else { + x_251 = x_249; +} +lean_ctor_set(x_251, 0, x_242); +lean_ctor_set(x_251, 1, x_243); +lean_ctor_set(x_251, 2, x_244); +lean_ctor_set(x_251, 3, x_245); +lean_ctor_set(x_251, 4, x_246); +lean_ctor_set(x_251, 5, x_250); +lean_ctor_set(x_251, 6, x_248); +x_252 = lean_st_ref_set(x_5, x_251, x_241); +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_254 = x_252; +} else { + lean_dec_ref(x_252); + x_254 = lean_box(0); +} +x_255 = lean_box(0); +if (lean_is_scalar(x_254)) { + x_256 = lean_alloc_ctor(0, 2, 0); +} else { + x_256 = x_254; +} +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_253); +return x_256; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__4(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 5); +lean_inc(x_6); +x_7 = l_Lean_logAt___at_Lean_addDecl___spec__5(x_6, x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_addDecl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2; +x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); +lean_dec(x_5); +if (x_7 == 0) +{ +uint8_t x_8; lean_object* x_9; +x_8 = 1; +x_9 = l_Lean_log___at_Lean_addDecl___spec__4(x_1, x_8, x_2, x_3, x_4); +return x_9; +} +else +{ +uint8_t x_10; lean_object* x_11; +x_10 = 2; +x_11 = l_Lean_log___at_Lean_addDecl___spec__4(x_1, x_10, x_2, x_3, x_4); +return x_11; +} +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = lean_ctor_get(x_1, 0); +lean_inc(x_5); +x_6 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = lean_ctor_get(x_1, 0); +lean_inc(x_7); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_4); +return x_8; +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_inc(x_7); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(x_1, x_5, x_2, x_3, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(x_4, x_7, x_8, x_11); +lean_dec(x_7); +lean_dec(x_4); +return x_12; +} +} +static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +double x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__1; +lean_inc(x_3); +lean_inc(x_1); +x_16 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_16, 0, x_1); +lean_ctor_set(x_16, 1, x_3); +lean_ctor_set_float(x_16, sizeof(void*)*2, x_15); +lean_ctor_set_float(x_16, sizeof(void*)*2 + 8, x_15); +lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 16, x_2); +x_17 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2; +x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_17); +lean_dec(x_7); +if (x_18 == 0) +{ +if (x_8 == 0) +{ +lean_object* x_19; lean_object* x_20; +lean_dec(x_3); +lean_dec(x_1); +x_19 = lean_box(0); +x_20 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); +return x_20; +} +else +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_16); +x_21 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_21, 0, x_1); +lean_ctor_set(x_21, 1, x_3); +lean_ctor_set_float(x_21, sizeof(void*)*2, x_9); +lean_ctor_set_float(x_21, sizeof(void*)*2 + 8, x_10); +lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 16, x_2); +x_22 = lean_box(0); +x_23 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); +return x_23; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_dec(x_16); +x_24 = lean_alloc_ctor(0, 2, 17); +lean_ctor_set(x_24, 0, x_1); +lean_ctor_set(x_24, 1, x_3); +lean_ctor_set_float(x_24, sizeof(void*)*2, x_9); +lean_ctor_set_float(x_24, sizeof(void*)*2 + 8, x_10); +lean_ctor_set_uint8(x_24, sizeof(void*)*2 + 16, x_2); +x_25 = lean_box(0); +x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); +return x_26; +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 53); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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, uint8_t x_7, double x_8, double x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +lean_object* x_15; lean_object* x_16; +lean_dec(x_11); +x_15 = lean_ctor_get(x_12, 5); +lean_inc(x_15); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_5); +x_16 = lean_apply_4(x_10, x_5, x_12, x_13, x_14); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_15, x_5, x_6, x_7, x_8, x_9, x_17, x_12, x_13, x_18); +lean_dec(x_13); +return x_19; +} +else +{ +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 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2; +x_22 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_15, x_5, x_6, x_7, x_8, x_9, x_21, x_12, x_13, x_20); +lean_dec(x_13); +return x_22; +} +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_useHearbeats; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_threshold; +return x_1; +} +} +static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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; uint8_t x_16; +lean_dec(x_8); +x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___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); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1; +x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_85 = lean_io_mono_nanos_now(x_14); +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +lean_inc(x_10); +lean_inc(x_9); +x_88 = lean_apply_3(x_7, x_9, x_10, x_87); +if (lean_obj_tag(x_88) == 0) +{ +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; double x_97; double x_98; double x_99; double x_100; double x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +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_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_89); +x_92 = lean_io_mono_nanos_now(x_90); +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 = 0; +x_96 = lean_unsigned_to_nat(0u); +x_97 = l_Float_ofScientific(x_86, x_95, x_96); +lean_dec(x_86); +x_98 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5; +x_99 = lean_float_div(x_97, x_98); +x_100 = l_Float_ofScientific(x_93, x_95, x_96); +lean_dec(x_93); +x_101 = lean_float_div(x_100, x_98); +x_102 = lean_box_float(x_99); +x_103 = lean_box_float(x_101); +x_104 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_104, 0, x_102); +lean_ctor_set(x_104, 1, x_103); +x_105 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_105, 0, x_91); +lean_ctor_set(x_105, 1, x_104); +x_17 = x_105; +x_18 = x_94; +goto block_84; +} +else +{ +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; double x_114; double x_115; double x_116; double x_117; double x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_106 = lean_ctor_get(x_88, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_88, 1); +lean_inc(x_107); +lean_dec(x_88); +x_108 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_108, 0, x_106); +x_109 = lean_io_mono_nanos_now(x_107); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = 0; +x_113 = lean_unsigned_to_nat(0u); +x_114 = l_Float_ofScientific(x_86, x_112, x_113); +lean_dec(x_86); +x_115 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5; +x_116 = lean_float_div(x_114, x_115); +x_117 = l_Float_ofScientific(x_110, x_112, x_113); +lean_dec(x_110); +x_118 = lean_float_div(x_117, x_115); +x_119 = lean_box_float(x_116); +x_120 = lean_box_float(x_118); +x_121 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_122, 0, x_108); +lean_ctor_set(x_122, 1, x_121); +x_17 = x_122; +x_18 = x_111; +goto block_84; +} +block_84: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_70; uint8_t x_71; +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 0); +lean_inc(x_20); +lean_dec(x_17); +x_21 = lean_ctor_get(x_19, 0); +lean_inc(x_21); +x_22 = lean_ctor_get(x_19, 1); +lean_inc(x_22); +lean_dec(x_19); +x_70 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2; +x_71 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_70); +if (x_71 == 0) +{ +uint8_t x_72; +x_72 = 0; +x_23 = x_72; +goto block_69; +} +else +{ +double x_73; double x_74; double x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; double x_80; double x_81; double x_82; uint8_t x_83; +x_73 = lean_unbox_float(x_22); +x_74 = lean_unbox_float(x_21); +x_75 = lean_float_sub(x_73, x_74); +x_76 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3; +x_77 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_76); +x_78 = 0; +x_79 = lean_unsigned_to_nat(0u); +x_80 = l_Float_ofScientific(x_77, x_78, x_79); +lean_dec(x_77); +x_81 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__4; +x_82 = lean_float_div(x_80, x_81); +x_83 = lean_float_decLt(x_82, x_75); +x_23 = x_83; +goto block_69; +} +block_69: +{ +if (x_6 == 0) +{ +if (x_23 == 0) +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_dec(x_22); +lean_dec(x_21); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_24 = lean_st_ref_take(x_10, x_18); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +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, 3); +x_29 = l_Lean_PersistentArray_append___rarg(x_13, x_28); +lean_ctor_set(x_25, 3, x_29); +x_30 = lean_st_ref_set(x_10, 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_addDecl___spec__7(x_20, x_9, x_10, x_31); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_20); +if (lean_obj_tag(x_32) == 0) +{ +uint8_t x_33; +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; +x_37 = !lean_is_exclusive(x_32); +if (x_37 == 0) +{ +return x_32; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +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_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +else +{ +lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_41 = lean_ctor_get(x_25, 0); +x_42 = lean_ctor_get(x_25, 1); +x_43 = lean_ctor_get(x_25, 2); +x_44 = lean_ctor_get(x_25, 3); +x_45 = lean_ctor_get(x_25, 4); +x_46 = lean_ctor_get(x_25, 5); +x_47 = lean_ctor_get(x_25, 6); +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_inc(x_41); +lean_dec(x_25); +x_48 = l_Lean_PersistentArray_append___rarg(x_13, x_44); +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_45); +lean_ctor_set(x_49, 5, x_46); +lean_ctor_set(x_49, 6, x_47); +x_50 = lean_st_ref_set(x_10, 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_addDecl___spec__7(x_20, x_9, x_10, x_51); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_20); +if (lean_obj_tag(x_52) == 0) +{ +lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +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_52); + x_55 = lean_box(0); +} +if (lean_is_scalar(x_55)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_55; +} +lean_ctor_set(x_56, 0, x_53); +lean_ctor_set(x_56, 1, x_54); +return x_56; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_57 = lean_ctor_get(x_52, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_52, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_52)) { + lean_ctor_release(x_52, 0); + lean_ctor_release(x_52, 1); + x_59 = x_52; +} else { + lean_dec_ref(x_52); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(1, 2, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +return x_60; +} +} +} +else +{ +lean_object* x_61; double x_62; double x_63; lean_object* x_64; +x_61 = lean_box(0); +x_62 = lean_unbox_float(x_21); +lean_dec(x_21); +x_63 = lean_unbox_float(x_22); +lean_dec(x_22); +x_64 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(x_2, x_3, x_4, x_13, x_20, x_1, x_23, x_62, x_63, x_5, x_61, x_9, x_10, x_18); +return x_64; +} +} +else +{ +lean_object* x_65; double x_66; double x_67; lean_object* x_68; +x_65 = lean_box(0); +x_66 = lean_unbox_float(x_21); +lean_dec(x_21); +x_67 = lean_unbox_float(x_22); +lean_dec(x_22); +x_68 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(x_2, x_3, x_4, x_13, x_20, x_1, x_23, x_66, x_67, x_5, x_65, x_9, x_10, x_18); +return x_68; +} +} +} +} +else +{ +lean_object* x_123; lean_object* x_124; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_189 = lean_io_get_num_heartbeats(x_14); +x_190 = lean_ctor_get(x_189, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); +lean_inc(x_10); +lean_inc(x_9); +x_192 = lean_apply_3(x_7, x_9, x_10, x_191); +if (lean_obj_tag(x_192) == 0) +{ +lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_object* x_200; double x_201; double x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_195, 0, x_193); +x_196 = lean_io_get_num_heartbeats(x_194); +x_197 = lean_ctor_get(x_196, 0); +lean_inc(x_197); +x_198 = lean_ctor_get(x_196, 1); +lean_inc(x_198); +lean_dec(x_196); +x_199 = 0; +x_200 = lean_unsigned_to_nat(0u); +x_201 = l_Float_ofScientific(x_190, x_199, x_200); +lean_dec(x_190); +x_202 = l_Float_ofScientific(x_197, x_199, x_200); +lean_dec(x_197); +x_203 = lean_box_float(x_201); +x_204 = lean_box_float(x_202); +x_205 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_206, 0, x_195); +lean_ctor_set(x_206, 1, x_205); +x_123 = x_206; +x_124 = x_198; +goto block_188; +} +else +{ +lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; double x_215; double x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_207 = lean_ctor_get(x_192, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_192, 1); +lean_inc(x_208); +lean_dec(x_192); +x_209 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_209, 0, x_207); +x_210 = lean_io_get_num_heartbeats(x_208); +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_210, 1); +lean_inc(x_212); +lean_dec(x_210); +x_213 = 0; +x_214 = lean_unsigned_to_nat(0u); +x_215 = l_Float_ofScientific(x_190, x_213, x_214); +lean_dec(x_190); +x_216 = l_Float_ofScientific(x_211, x_213, x_214); +lean_dec(x_211); +x_217 = lean_box_float(x_215); +x_218 = lean_box_float(x_216); +x_219 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_219, 0, x_217); +lean_ctor_set(x_219, 1, x_218); +x_220 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_220, 0, x_209); +lean_ctor_set(x_220, 1, x_219); +x_123 = x_220; +x_124 = x_212; +goto block_188; +} +block_188: +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_176; uint8_t x_177; +x_125 = lean_ctor_get(x_123, 1); +lean_inc(x_125); +x_126 = lean_ctor_get(x_123, 0); +lean_inc(x_126); +lean_dec(x_123); +x_127 = lean_ctor_get(x_125, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_125, 1); +lean_inc(x_128); +lean_dec(x_125); +x_176 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2; +x_177 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_176); +if (x_177 == 0) +{ +uint8_t x_178; +x_178 = 0; +x_129 = x_178; +goto block_175; +} +else +{ +double x_179; double x_180; double x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; lean_object* x_185; double x_186; uint8_t x_187; +x_179 = lean_unbox_float(x_128); +x_180 = lean_unbox_float(x_127); +x_181 = lean_float_sub(x_179, x_180); +x_182 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3; +x_183 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_182); +x_184 = 0; +x_185 = lean_unsigned_to_nat(0u); +x_186 = l_Float_ofScientific(x_183, x_184, x_185); +lean_dec(x_183); +x_187 = lean_float_decLt(x_186, x_181); +x_129 = x_187; +goto block_175; +} +block_175: +{ +if (x_6 == 0) +{ +if (x_129 == 0) +{ +lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; +lean_dec(x_128); +lean_dec(x_127); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_130 = lean_st_ref_take(x_10, x_124); +x_131 = lean_ctor_get(x_130, 0); +lean_inc(x_131); +x_132 = lean_ctor_get(x_130, 1); +lean_inc(x_132); +lean_dec(x_130); +x_133 = !lean_is_exclusive(x_131); +if (x_133 == 0) +{ +lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; +x_134 = lean_ctor_get(x_131, 3); +x_135 = l_Lean_PersistentArray_append___rarg(x_13, x_134); +lean_ctor_set(x_131, 3, x_135); +x_136 = lean_st_ref_set(x_10, x_131, x_132); +x_137 = lean_ctor_get(x_136, 1); +lean_inc(x_137); +lean_dec(x_136); +x_138 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(x_126, x_9, x_10, x_137); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_126); +if (lean_obj_tag(x_138) == 0) +{ +uint8_t x_139; +x_139 = !lean_is_exclusive(x_138); +if (x_139 == 0) +{ +return x_138; +} +else +{ +lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_140 = lean_ctor_get(x_138, 0); +x_141 = lean_ctor_get(x_138, 1); +lean_inc(x_141); +lean_inc(x_140); +lean_dec(x_138); +x_142 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_142, 0, x_140); +lean_ctor_set(x_142, 1, x_141); +return x_142; +} +} +else +{ +uint8_t x_143; +x_143 = !lean_is_exclusive(x_138); +if (x_143 == 0) +{ +return x_138; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_138, 0); +x_145 = lean_ctor_get(x_138, 1); +lean_inc(x_145); +lean_inc(x_144); +lean_dec(x_138); +x_146 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_146, 0, x_144); +lean_ctor_set(x_146, 1, x_145); +return x_146; +} +} +} +else +{ +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_147 = lean_ctor_get(x_131, 0); +x_148 = lean_ctor_get(x_131, 1); +x_149 = lean_ctor_get(x_131, 2); +x_150 = lean_ctor_get(x_131, 3); +x_151 = lean_ctor_get(x_131, 4); +x_152 = lean_ctor_get(x_131, 5); +x_153 = lean_ctor_get(x_131, 6); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_131); +x_154 = l_Lean_PersistentArray_append___rarg(x_13, x_150); +x_155 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_155, 0, x_147); +lean_ctor_set(x_155, 1, x_148); +lean_ctor_set(x_155, 2, x_149); +lean_ctor_set(x_155, 3, x_154); +lean_ctor_set(x_155, 4, x_151); +lean_ctor_set(x_155, 5, x_152); +lean_ctor_set(x_155, 6, x_153); +x_156 = lean_st_ref_set(x_10, x_155, x_132); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +lean_dec(x_156); +x_158 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(x_126, x_9, x_10, x_157); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_126); +if (lean_obj_tag(x_158) == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; +x_159 = lean_ctor_get(x_158, 0); +lean_inc(x_159); +x_160 = lean_ctor_get(x_158, 1); +lean_inc(x_160); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_161 = x_158; +} else { + lean_dec_ref(x_158); + x_161 = lean_box(0); +} +if (lean_is_scalar(x_161)) { + x_162 = lean_alloc_ctor(0, 2, 0); +} else { + x_162 = x_161; +} +lean_ctor_set(x_162, 0, x_159); +lean_ctor_set(x_162, 1, x_160); +return x_162; +} +else +{ +lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_163 = lean_ctor_get(x_158, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_158, 1); +lean_inc(x_164); +if (lean_is_exclusive(x_158)) { + lean_ctor_release(x_158, 0); + lean_ctor_release(x_158, 1); + x_165 = x_158; +} else { + lean_dec_ref(x_158); + x_165 = lean_box(0); +} +if (lean_is_scalar(x_165)) { + x_166 = lean_alloc_ctor(1, 2, 0); +} else { + x_166 = x_165; +} +lean_ctor_set(x_166, 0, x_163); +lean_ctor_set(x_166, 1, x_164); +return x_166; +} +} +} +else +{ +lean_object* x_167; double x_168; double x_169; lean_object* x_170; +x_167 = lean_box(0); +x_168 = lean_unbox_float(x_127); +lean_dec(x_127); +x_169 = lean_unbox_float(x_128); +lean_dec(x_128); +x_170 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(x_2, x_3, x_4, x_13, x_126, x_1, x_129, x_168, x_169, x_5, x_167, x_9, x_10, x_124); +return x_170; +} +} +else +{ +lean_object* x_171; double x_172; double x_173; lean_object* x_174; +x_171 = lean_box(0); +x_172 = lean_unbox_float(x_127); +lean_dec(x_127); +x_173 = lean_unbox_float(x_128); +lean_dec(x_128); +x_174 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(x_2, x_3, x_4, x_13, x_126, x_1, x_129, x_172, x_173, x_5, x_171, x_9, x_10, x_124); +return x_174; +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; +x_9 = lean_ctor_get(x_6, 2); +lean_inc(x_9); +lean_inc(x_1); +x_10 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(x_1, x_6, x_7, x_8); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_unbox(x_11); +if (x_12 == 0) +{ +lean_object* x_13; lean_object* x_14; uint8_t x_15; +x_13 = lean_ctor_get(x_10, 1); +lean_inc(x_13); +lean_dec(x_10); +x_14 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2; +x_15 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_14); +if (x_15 == 0) +{ +lean_object* x_16; +lean_dec(x_11); +lean_dec(x_9); +lean_dec(x_5); +lean_dec(x_2); +lean_dec(x_1); +x_16 = lean_apply_3(x_3, x_6, x_7, x_13); +if (lean_obj_tag(x_16) == 0) +{ +uint8_t x_17; +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) +{ +return x_16; +} +else +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 0); +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_19); +return x_20; +} +} +else +{ +uint8_t x_21; +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) +{ +return x_16; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; +} +} +} +else +{ +lean_object* x_25; uint8_t x_26; lean_object* x_27; +x_25 = lean_box(0); +x_26 = lean_unbox(x_11); +lean_dec(x_11); +x_27 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4(x_9, x_1, x_4, x_5, x_2, x_26, x_3, x_25, x_6, x_7, x_13); +return x_27; +} +} +else +{ +lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; +x_28 = lean_ctor_get(x_10, 1); +lean_inc(x_28); +lean_dec(x_10); +x_29 = lean_box(0); +x_30 = lean_unbox(x_11); +lean_dec(x_11); +x_31 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4(x_9, x_1, x_4, x_5, x_2, x_30, x_3, x_29, x_6, x_7, x_28); +return x_31; +} +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_apply_2(x_3, x_5, x_6); +x_9 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_8, x_4, x_7); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg___boxed), 7, 0); +return x_2; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("typechecking declaration", 24); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__1___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; +x_5 = l_Lean_addDecl___lambda__1___closed__2; +x_6 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_6, 0, x_5); +lean_ctor_set(x_6, 1, x_4); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_2); +x_6 = lean_st_ref_get(x_4, x_5); +x_7 = lean_ctor_get(x_6, 0); +lean_inc(x_7); +x_8 = lean_ctor_get(x_6, 1); +lean_inc(x_8); +lean_dec(x_6); +x_9 = lean_ctor_get(x_7, 0); +lean_inc(x_9); +lean_dec(x_7); +x_10 = lean_ctor_get(x_3, 2); +lean_inc(x_10); +x_11 = l_Lean_Environment_addDecl(x_9, x_10, x_1); +lean_dec(x_1); +lean_dec(x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_12, x_3, x_4, x_8); +lean_dec(x_4); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 0); +lean_inc(x_14); +lean_dec(x_11); +x_15 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_14, x_3, x_4, x_8); +lean_dec(x_4); +lean_dec(x_3); +return x_15; +} +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("declaration uses 'sorry'", 24); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__3___closed__1; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_addDecl___lambda__3___closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_addDecl___lambda__3___closed__2; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; +x_5 = lean_st_ref_get(x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = lean_ctor_get(x_6, 5); +lean_inc(x_8); +lean_dec(x_6); +x_9 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_8); +if (x_9 == 0) +{ +uint8_t x_10; uint8_t x_11; +x_10 = 0; +lean_inc(x_1); +x_11 = l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(x_1, x_10); +if (x_11 == 0) +{ +lean_object* x_12; lean_object* x_13; +x_12 = lean_box(0); +x_13 = l_Lean_addDecl___lambda__2(x_1, x_12, x_2, x_3, x_7); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_14 = l_Lean_addDecl___lambda__3___closed__3; +lean_inc(x_3); +lean_inc(x_2); +x_15 = l_Lean_logWarning___at_Lean_addDecl___spec__3(x_14, x_2, x_3, x_7); +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_addDecl___lambda__2(x_1, x_16, x_2, x_3, x_17); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = lean_box(0); +x_20 = l_Lean_addDecl___lambda__2(x_1, x_19, x_2, x_3, x_7); +return x_20; +} +} +} +static lean_object* _init_l_Lean_addDecl___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Kernel", 6); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Lean_addDecl___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_addDecl___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_addDecl___closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("type checking", 13); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; 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; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__3), 4, 1); +lean_closure_set(x_6, 0, x_1); +x_7 = l_Lean_addDecl___closed__2; +x_8 = l_Lean_addDecl___closed__3; +x_9 = 1; +x_10 = l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1; +x_11 = lean_box(x_9); +x_12 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___boxed), 8, 5); +lean_closure_set(x_12, 0, x_7); +lean_closure_set(x_12, 1, x_8); +lean_closure_set(x_12, 2, x_6); +lean_closure_set(x_12, 3, x_11); +lean_closure_set(x_12, 4, x_10); +x_13 = l_Lean_addDecl___closed__4; +x_14 = lean_box(0); +x_15 = l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(x_13, x_5, x_12, x_14, x_2, x_3, x_4); +lean_dec(x_5); +return x_15; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___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: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_logAt___at_Lean_addDecl___spec__5(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__4___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_6; lean_object* x_7; +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = l_Lean_log___at_Lean_addDecl___spec__4(x_1, x_6, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__7(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_6); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; double x_18; lean_object* x_19; +x_15 = lean_unbox(x_2); +lean_dec(x_2); +x_16 = lean_unbox(x_8); +lean_dec(x_8); +x_17 = lean_unbox_float(x_9); +lean_dec(x_9); +x_18 = lean_unbox_float(x_10); +lean_dec(x_10); +x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_16, x_17, x_18, x_11, x_12, x_13, x_14); +lean_dec(x_13); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { +_start: +{ +uint8_t x_15; uint8_t x_16; double x_17; double x_18; lean_object* x_19; +x_15 = lean_unbox(x_2); +lean_dec(x_2); +x_16 = lean_unbox(x_7); +lean_dec(x_7); +x_17 = lean_unbox_float(x_8); +lean_dec(x_8); +x_18 = lean_unbox_float(x_9); +lean_dec(x_9); +x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3(x_1, x_15, x_3, x_4, x_5, x_6, x_16, x_17, x_18, x_10, x_11, x_12, x_13, x_14); +return x_19; +} +} +LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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; +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_addDecl___spec__6___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_addDecl___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +uint8_t x_9; lean_object* x_10; +x_9 = lean_unbox(x_4); +lean_dec(x_4); +x_10 = l_Lean_withTraceNode___at_Lean_addDecl___spec__6(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); +return x_10; +} +} +LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___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__8___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_2); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_addDecl___lambda__1(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +lean_inc(x_3); +lean_inc(x_2); +lean_inc(x_1); +x_5 = l_Lean_addDecl(x_1, x_2, x_3, x_4); +if (lean_obj_tag(x_5) == 0) +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_5, 1); +lean_inc(x_6); +lean_dec(x_5); +x_7 = l_Lean_compileDecl(x_1, x_2, x_3, x_6); +return x_7; +} +else +{ +uint8_t x_8; +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_8 = !lean_is_exclusive(x_5); +if (x_8 == 0) +{ +return x_5; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = lean_ctor_get(x_5, 0); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_5); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_9); +lean_ctor_set(x_11, 1, x_10); +return x_11; +} +} +} +} +lean_object* initialize_Lean_CoreM(uint8_t builtin, lean_object*); +static bool _G_initialized = false; +LEAN_EXPORT lean_object* initialize_Lean_AddDecl(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_CoreM(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); +l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1 = _init_l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1(); +lean_mark_persistent(l_Lean_logAt___at_Lean_addDecl___spec__5___closed__1); +l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2 = _init_l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2(); +lean_mark_persistent(l_Lean_logAt___at_Lean_addDecl___spec__5___closed__2); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__1(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__2___closed__2); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__1); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__2); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__3); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__4 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__4(); +l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__6___lambda__4___closed__5(); +l_Lean_addDecl___lambda__1___closed__1 = _init_l_Lean_addDecl___lambda__1___closed__1(); +lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__1); +l_Lean_addDecl___lambda__1___closed__2 = _init_l_Lean_addDecl___lambda__1___closed__2(); +lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__2); +l_Lean_addDecl___lambda__3___closed__1 = _init_l_Lean_addDecl___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__1); +l_Lean_addDecl___lambda__3___closed__2 = _init_l_Lean_addDecl___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__2); +l_Lean_addDecl___lambda__3___closed__3 = _init_l_Lean_addDecl___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__3); +l_Lean_addDecl___closed__1 = _init_l_Lean_addDecl___closed__1(); +lean_mark_persistent(l_Lean_addDecl___closed__1); +l_Lean_addDecl___closed__2 = _init_l_Lean_addDecl___closed__2(); +lean_mark_persistent(l_Lean_addDecl___closed__2); +l_Lean_addDecl___closed__3 = _init_l_Lean_addDecl___closed__3(); +lean_mark_persistent(l_Lean_addDecl___closed__3); +l_Lean_addDecl___closed__4 = _init_l_Lean_addDecl___closed__4(); +lean_mark_persistent(l_Lean_addDecl___closed__4); +return lean_io_result_mk_ok(lean_box(0)); +} +#ifdef __cplusplus +} +#endif diff --git a/stage0/stdlib/Lean/Attributes.c b/stage0/stdlib/Lean/Attributes.c index 2f76319eceb1..6724954261e5 100644 --- a/stage0/stdlib/Lean/Attributes.c +++ b/stage0/stdlib/Lean/Attributes.c @@ -21,7 +21,6 @@ LEAN_EXPORT lean_object* l_Lean_ParametricAttributeImpl_afterImport___default___ LEAN_EXPORT uint8_t l_Lean_instInhabitedAttributeApplicationTime; LEAN_EXPORT lean_object* l_Lean_instMonadLiftImportMAttrM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getAttrParamOptPrio(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerBuiltinAttribute___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerParametricAttribute___rarg___closed__1; static lean_object* l_Lean_AttributeImpl_erase___default___rarg___closed__2; @@ -36,6 +35,7 @@ static lean_object* l_Lean_EnumAttributes_setValue___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_getId(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_getBuiltinAttributeImpl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__11; +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3114____spec__1___boxed(lean_object*); static lean_object* l_Lean_registerParametricAttribute___rarg___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_getPrio(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_mkAttributeImplOfBuilder___spec__1(lean_object*, lean_object*); @@ -43,6 +43,7 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedAttributeExtensionState; static lean_object* l_Lean_registerEnumAttributes___rarg___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_registerEnumAttributes___rarg___lambda__2(lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2___boxed(lean_object*); static lean_object* l_Lean_registerBuiltinAttribute___lambda__1___closed__1; uint8_t lean_usize_dec_le(size_t, size_t); static lean_object* l_Lean_registerParametricAttribute___rarg___closed__4; @@ -50,16 +51,15 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_updateEnvAttr LEAN_EXPORT lean_object* l_Lean_registerEnumAttributes___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getAttributeImpl___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___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_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_292_(lean_object*); LEAN_EXPORT lean_object* l_Lean_getAttributeNames___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerAttributeImplBuilder___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Attributes_0__Lean_AttributeExtension_addImported___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_getBuiltinAttributeImpl___spec__3(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_instMonadLiftImportMAttrM(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2; lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_Lean_instBEqAttributeKind___closed__1; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -87,18 +87,19 @@ static lean_object* l_Lean_registerTagAttribute___lambda__6___closed__3; static lean_object* l_Lean_instInhabitedTagAttribute___closed__6; static lean_object* l_Lean_isAttribute___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_registerBuiltinAttribute___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Attributes___hyg_1253_; LEAN_EXPORT lean_object* l_Lean_EnumAttributes_getValue(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_registerBuiltinAttribute___spec__5(lean_object*, lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, 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_Lean_initFn____x40_Lean_Attributes___hyg_3114_(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_ParametricAttribute_setParam___spec__1___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_getBuiltinAttributeNames___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718_(lean_object*); static lean_object* l_Lean_EnumAttributes_setValue___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Attribute_add(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Attributes___hyg_1283_; static lean_object* l_Lean_registerAttributeImplBuilder___closed__1; static lean_object* l_Lean_instInhabitedParametricAttribute___closed__1; uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); @@ -122,11 +123,10 @@ lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1(lea static lean_object* l_Lean_registerTagAttribute___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_ParametricAttribute_getParam_x3f___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_EnumAttributes_setValue___spec__1___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7; static lean_object* l_Lean_instInhabitedAttributeImpl___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_registerAttributeImplBuilder___spec__3(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_isAttribute___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2(lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__27; static lean_object* l_Lean_registerTagAttribute___closed__1; lean_object* l_Lean_initializing(lean_object*); @@ -138,6 +138,7 @@ LEAN_EXPORT lean_object* l_Lean_registerTagAttribute___lambda__5___boxed(lean_ob LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_getIdent___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Attribute_add___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerTagAttribute___lambda__4___closed__5; +LEAN_EXPORT uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_registerEnumAttributes___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Attribute_Builtin_ensureNoArgs___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); @@ -151,7 +152,6 @@ LEAN_EXPORT lean_object* l_Lean_AttributeApplicationTime_noConfusion___rarg___la lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_registerEnumAttributes___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_getAttrParamOptPrio___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4; static lean_object* l_Lean_registerBuiltinAttribute___closed__2; LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___rarg___lambda__5(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*); @@ -171,27 +171,25 @@ LEAN_EXPORT uint8_t l_Lean_TagAttribute_hasTag(lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_registerAttributeImplBuilder___spec__7___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_registerAttributeImplBuilder___spec__2___boxed(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getBuiltinAttributeImpl___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Attribute_Builtin_getIdent_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1; lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AttributeKind_noConfusion___rarg(uint8_t, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkAttributeImplOfEntry(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__19; lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3204____spec__1___boxed(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5; LEAN_EXPORT lean_object* l_Lean_instInhabitedAttributeImpl___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTagAttribute___closed__4; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_registerAttributeImplBuilder___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AttributeApplicationTime_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_instToStringAttributeKind___closed__1; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3204_(lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringAttributeKind(uint8_t); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerTagAttribute___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerParametricAttribute___rarg___closed__2; -LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_registerParametricAttribute___spec__2___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__28; LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_registerParametricAttribute___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -220,6 +218,7 @@ LEAN_EXPORT lean_object* l_Lean_AttributeExtensionState_newEntries___default; static lean_object* l_Lean_registerTagAttribute___lambda__6___closed__1; static lean_object* l_Lean_instInhabitedAttributeImpl___closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_registerEnumAttributes___spec__3___rarg___lambda__1___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_Attributes___hyg_262____closed__3; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__15; LEAN_EXPORT lean_object* l_Lean_getBuiltinAttributeImpl(lean_object*, lean_object*); @@ -236,6 +235,7 @@ static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_beqAttributeApplicationTime____x40_Lean_Attributes___hyg_15____boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__29; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Attributes___hyg_2486_; LEAN_EXPORT lean_object* l_Lean_instInhabitedAttributeImpl___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_ParametricAttribute_getParam_x3f___spec__2(lean_object*); static lean_object* l_Lean_registerBuiltinAttribute___lambda__2___closed__1; @@ -259,15 +259,16 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedTagAttribute___lambda__2(lean_objec LEAN_EXPORT lean_object* l_Lean_registerAttributeImplBuilder___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_attribute_application_time(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3114____spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_addAttrEntry(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTagAttribute___lambda__3___boxed(lean_object*); uint8_t l_Lean_Syntax_isIdent(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5; LEAN_EXPORT lean_object* lean_is_attribute(lean_object*, lean_object*); static lean_object* l_Lean_registerTagAttribute___lambda__6___closed__2; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_mkAttributeImplOfBuilder___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_registerEnumAttributes___spec__1___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_registerBuiltinAttribute___lambda__2___closed__2; -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_registerBuiltinAttribute___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_registerTagAttribute(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerBuiltinAttribute___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -280,15 +281,16 @@ static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__16; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_getIdent_x3f___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnumAttributes_getValue___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instBEqAttributeKind; +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2; lean_object* lean_usize_to_nat(size_t); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_getBuiltinAttributeImpl___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_getBuiltinAttributeNames(lean_object*); LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_getId_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_mkAttributeImplOfBuilder___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153____boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at_Lean_registerParametricAttribute___spec__2___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___rarg___lambda__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ParametricAttribute_setParam___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -303,10 +305,12 @@ static lean_object* l_Lean_registerAttributeImplBuilder___closed__2; LEAN_EXPORT lean_object* l_Lean_registerTagAttribute___lambda__4(lean_object*); static lean_object* l_Lean_registerTagAttribute___lambda__4___closed__4; LEAN_EXPORT lean_object* l_List_forM___at_Lean_registerEnumAttributes___spec__4(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__1(lean_object*); static lean_object* l_Lean_instInhabitedAttributeImpl___closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_262_(lean_object*); LEAN_EXPORT lean_object* l_Lean_ParametricAttribute_getParam_x3f(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3; LEAN_EXPORT lean_object* l_Lean_AttributeApplicationTime_toCtorIdx___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Attributes_0__Lean_AttributeExtension_addImported___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ParametricAttribute_setParam___rarg___closed__3; @@ -329,9 +333,7 @@ LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___boxed(lean_object* LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_EnumAttributes_getValue___spec__1___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__21; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_registerEnumAttributes___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_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_AttributeImpl_erase___default___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_registerAttributeImplBuilder___spec__5(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -343,7 +345,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_registerEnumAttributes___spec__1(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_AttributeExtension_mkInitial(lean_object*); static lean_object* l_Lean_getAttrParamOptPrio___closed__1; -LEAN_EXPORT uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__4; static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__1; LEAN_EXPORT lean_object* l_Lean_registerEnumAttributes___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); @@ -358,6 +359,7 @@ LEAN_EXPORT lean_object* l_Lean_registerTagAttribute___boxed(lean_object*, lean_ static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__25; static lean_object* l_Lean_registerEnumAttributes___rarg___lambda__3___closed__2; static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__20; +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1; uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_ParametricAttribute_getParam_x3f___spec__1___rarg___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__8; @@ -370,7 +372,6 @@ uint32_t lean_uint32_of_nat(lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__9; LEAN_EXPORT lean_object* l_Lean_registerAttributeOfDecl___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3204____spec__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_registerBuiltinAttribute___spec__6(lean_object*, size_t, lean_object*); static lean_object* l_Lean_mkAttributeImplOfConstantUnsafe___closed__1; static lean_object* l_Lean_instInhabitedTagAttribute___closed__1; @@ -402,7 +403,6 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Attribute LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_EnumAttributes_getValue___spec__2(lean_object*); static lean_object* l_Lean_registerAttributeImplBuilder___lambda__1___closed__1; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_registerAttributeImplBuilder___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTagAttribute___lambda__4___boxed(lean_object*); static lean_object* l_Lean_registerTagAttribute___lambda__7___closed__1; @@ -431,6 +431,7 @@ static lean_object* l_Lean_instInhabitedTagAttribute___closed__2; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_registerEnumAttributes___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedTagAttribute___closed__7; +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3; LEAN_EXPORT lean_object* l_Lean_getAttributeImpl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedTagAttribute___lambda__4(lean_object*); static lean_object* l_Lean_Attribute_Builtin_ensureNoArgs___closed__7; @@ -458,7 +459,6 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_registerAttributeImplBu LEAN_EXPORT lean_object* l_Lean_instInhabitedEnumAttributes(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1; LEAN_EXPORT lean_object* l_Lean_instBEqAttributeApplicationTime; uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___rarg___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -472,7 +472,9 @@ static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__24; static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__14; static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_registerBuiltinAttribute___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_registerBuiltinAttribute___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6; LEAN_EXPORT lean_object* l_Lean_Attribute_Builtin_getIdent_x3f(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerTagAttribute___lambda__8(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); @@ -482,7 +484,6 @@ LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_registerEnumAttributes___s lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ParametricAttributeImpl_afterSet___default___rarg(lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Attributes___hyg_2546_; LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___rarg___lambda__3(lean_object*); static lean_object* l___auto____x40_Lean_Attributes___hyg_98____closed__17; static lean_object* l_Lean_instInhabitedEnumAttributes___closed__1; @@ -494,7 +495,6 @@ LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_registerAttributeImplBu LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Attribute_Builtin_getIdent___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_attributeImplBuilderTableRef; LEAN_EXPORT lean_object* l_Lean_instInhabitedAttributeImplCore; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808_(lean_object*); size_t lean_usize_land(size_t, size_t); lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); @@ -1158,7 +1158,7 @@ x_6 = l_Lean_AttributeKind_noConfusion___rarg(x_4, x_5, x_3); return x_6; } } -LEAN_EXPORT uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t x_1, uint8_t x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t x_1, uint8_t x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -1170,7 +1170,7 @@ lean_dec(x_3); return x_5; } } -LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153____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; @@ -1178,7 +1178,7 @@ x_3 = lean_unbox(x_1); lean_dec(x_1); x_4 = lean_unbox(x_2); lean_dec(x_2); -x_5 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_4); +x_5 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_4); x_6 = lean_box(x_5); return x_6; } @@ -1187,7 +1187,7 @@ static lean_object* _init_l_Lean_instBEqAttributeKind___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153____boxed), 2, 0); return x_1; } } @@ -1466,7 +1466,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1() { _start: { lean_object* x_1; @@ -1474,21 +1474,21 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1; +x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2; +x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -1496,11 +1496,11 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_292_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_262_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3; x_3 = lean_st_mk_ref(x_2, x_1); x_4 = !lean_is_exclusive(x_3); if (x_4 == 0) @@ -3875,7 +3875,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Attributes___hyg_1283_() { +static lean_object* _init_l___auto____x40_Lean_Attributes___hyg_1253_() { _start: { lean_object* x_1; @@ -4037,7 +4037,7 @@ static lean_object* _init_l_Lean_registerTagAttribute___lambda__5___closed__1() _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3; x_2 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_2, 0, x_1); lean_ctor_set(x_2, 1, x_1); @@ -4304,7 +4304,7 @@ x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); x_12 = 0; -x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_6, x_12); +x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_6, 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; uint8_t x_20; @@ -5465,7 +5465,7 @@ LEAN_EXPORT lean_object* l_Lean_registerParametricAttribute___rarg___lambda__7(l { uint8_t x_11; uint8_t x_12; x_11 = 0; -x_12 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_7, x_11); +x_12 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_7, x_11); if (x_12 == 0) { lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; @@ -6235,7 +6235,7 @@ x_2 = l_Lean_instInhabitedEnumAttributes___closed__1; return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Attributes___hyg_2546_() { +static lean_object* _init_l___auto____x40_Lean_Attributes___hyg_2486_() { _start: { lean_object* x_1; @@ -6546,7 +6546,7 @@ x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); x_13 = 0; -x_14 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_7, x_13); +x_14 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_7, 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; uint8_t x_21; @@ -7532,7 +7532,7 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3204____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3114____spec__1(lean_object* x_1) { _start: { lean_object* x_2; @@ -7540,7 +7540,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3204_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3114_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -7567,11 +7567,11 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3204____spec__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3114____spec__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3204____spec__1(x_1); +x_2 = l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_Attributes___hyg_3114____spec__1(x_1); lean_dec(x_1); return x_2; } @@ -8289,7 +8289,7 @@ static lean_object* _init_l_Lean_instInhabitedAttributeExtensionState___closed__ { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); lean_ctor_set(x_3, 1, x_2); @@ -9054,7 +9054,7 @@ lean_ctor_set(x_11, 1, x_10); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__1(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; @@ -9069,7 +9069,7 @@ x_6 = l_List_toArrayAux___rarg(x_3, x_5); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2(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; @@ -9086,7 +9086,7 @@ lean_ctor_set(x_8, 1, x_6); return x_8; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1() { _start: { lean_object* x_1; @@ -9094,17 +9094,17 @@ x_1 = lean_mk_string_from_bytes("attributeExtension", 18); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___auto____x40_Lean_Attributes___hyg_98____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3() { _start: { lean_object* x_1; @@ -9112,7 +9112,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Attributes_0__Lean_AttributeEx return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4() { _start: { lean_object* x_1; @@ -9120,7 +9120,7 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Attributes_0__Lean_AttributeEx return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5() { _start: { lean_object* x_1; @@ -9128,32 +9128,32 @@ x_1 = lean_alloc_closure((void*)(l___private_Lean_Attributes_0__Lean_addAttrEntr return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__1), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__1), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2___boxed), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3; -x_3 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4; -x_4 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5; -x_5 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6; -x_6 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7; +x_1 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3; +x_3 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5; +x_5 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6; +x_6 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7; x_7 = lean_alloc_ctor(0, 6, 0); lean_ctor_set(x_7, 0, x_1); lean_ctor_set(x_7, 1, x_2); @@ -9164,20 +9164,20 @@ lean_ctor_set(x_7, 5, x_6); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8; +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8; x_3 = l_Lean_registerPersistentEnvExtensionUnsafe___rarg(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3808____lambda__2(x_1); +x_2 = l_Lean_initFn____x40_Lean_Attributes___hyg_3718____lambda__2(x_1); lean_dec(x_1); return x_2; } @@ -10453,13 +10453,13 @@ l_Lean_instInhabitedAttributeImpl___closed__3 = _init_l_Lean_instInhabitedAttrib lean_mark_persistent(l_Lean_instInhabitedAttributeImpl___closed__3); l_Lean_instInhabitedAttributeImpl = _init_l_Lean_instInhabitedAttributeImpl(); lean_mark_persistent(l_Lean_instInhabitedAttributeImpl); -l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__1); -l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__2); -l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_292____closed__3); -if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_292_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__1); +l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__2); +l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_262____closed__3); +if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_262_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_attributeMapRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_attributeMapRef); @@ -10533,8 +10533,8 @@ l_Lean_instInhabitedTagAttribute___closed__8 = _init_l_Lean_instInhabitedTagAttr lean_mark_persistent(l_Lean_instInhabitedTagAttribute___closed__8); l_Lean_instInhabitedTagAttribute = _init_l_Lean_instInhabitedTagAttribute(); lean_mark_persistent(l_Lean_instInhabitedTagAttribute); -l___auto____x40_Lean_Attributes___hyg_1283_ = _init_l___auto____x40_Lean_Attributes___hyg_1283_(); -lean_mark_persistent(l___auto____x40_Lean_Attributes___hyg_1283_); +l___auto____x40_Lean_Attributes___hyg_1253_ = _init_l___auto____x40_Lean_Attributes___hyg_1253_(); +lean_mark_persistent(l___auto____x40_Lean_Attributes___hyg_1253_); l_Lean_registerTagAttribute___lambda__4___closed__1 = _init_l_Lean_registerTagAttribute___lambda__4___closed__1(); lean_mark_persistent(l_Lean_registerTagAttribute___lambda__4___closed__1); l_Lean_registerTagAttribute___lambda__4___closed__2 = _init_l_Lean_registerTagAttribute___lambda__4___closed__2(); @@ -10601,8 +10601,8 @@ l_Lean_ParametricAttribute_setParam___rarg___closed__3 = _init_l_Lean_Parametric lean_mark_persistent(l_Lean_ParametricAttribute_setParam___rarg___closed__3); l_Lean_instInhabitedEnumAttributes___closed__1 = _init_l_Lean_instInhabitedEnumAttributes___closed__1(); lean_mark_persistent(l_Lean_instInhabitedEnumAttributes___closed__1); -l___auto____x40_Lean_Attributes___hyg_2546_ = _init_l___auto____x40_Lean_Attributes___hyg_2546_(); -lean_mark_persistent(l___auto____x40_Lean_Attributes___hyg_2546_); +l___auto____x40_Lean_Attributes___hyg_2486_ = _init_l___auto____x40_Lean_Attributes___hyg_2486_(); +lean_mark_persistent(l___auto____x40_Lean_Attributes___hyg_2486_); l_Lean_registerEnumAttributes___rarg___lambda__3___closed__1 = _init_l_Lean_registerEnumAttributes___rarg___lambda__3___closed__1(); lean_mark_persistent(l_Lean_registerEnumAttributes___rarg___lambda__3___closed__1); l_Lean_registerEnumAttributes___rarg___lambda__3___closed__2 = _init_l_Lean_registerEnumAttributes___rarg___lambda__3___closed__2(); @@ -10619,7 +10619,7 @@ l_Lean_EnumAttributes_setValue___rarg___closed__1 = _init_l_Lean_EnumAttributes_ lean_mark_persistent(l_Lean_EnumAttributes_setValue___rarg___closed__1); l_Lean_EnumAttributes_setValue___rarg___closed__2 = _init_l_Lean_EnumAttributes_setValue___rarg___closed__2(); lean_mark_persistent(l_Lean_EnumAttributes_setValue___rarg___closed__2); -if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_3204_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_3114_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_attributeImplBuilderTableRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_attributeImplBuilderTableRef); @@ -10648,23 +10648,23 @@ l_Lean_mkAttributeImplOfConstantUnsafe___closed__3 = _init_l_Lean_mkAttributeImp lean_mark_persistent(l_Lean_mkAttributeImplOfConstantUnsafe___closed__3); l_Lean_mkAttributeImplOfConstantUnsafe___closed__4 = _init_l_Lean_mkAttributeImplOfConstantUnsafe___closed__4(); lean_mark_persistent(l_Lean_mkAttributeImplOfConstantUnsafe___closed__4); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__1); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__2); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__3); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__4); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__5); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__6); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__7); -l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3808____closed__8); -if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_3808_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__1); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__2); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__3); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__4); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__5); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__6); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__7); +l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8 = _init_l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Attributes___hyg_3718____closed__8); +if (builtin) {res = l_Lean_initFn____x40_Lean_Attributes___hyg_3718_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_attributeExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_attributeExtension); diff --git a/stage0/stdlib/Lean/Class.c b/stage0/stdlib/Lean/Class.c index 6087ebb400d0..f2bbf15518a6 100644 --- a/stage0/stdlib/Lean/Class.c +++ b/stage0/stdlib/Lean/Class.c @@ -71,6 +71,7 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedClassState; LEAN_EXPORT lean_object* l_Lean_ClassEntry_lt___boxed(lean_object*, lean_object*); size_t lean_ptr_addr(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Class___hyg_768____closed__18; +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); static lean_object* l_Lean_mkOutParamArgsImplicit_go___closed__6; static lean_object* l_Lean_SMap_empty___at_Lean_ClassState_outParamMap___default___spec__1___closed__3; size_t lean_usize_of_nat(lean_object*); @@ -169,7 +170,6 @@ static lean_object* l_Lean_initFn____x40_Lean_Class___hyg_768____lambda__3___clo static lean_object* l_Lean_addClass___lambda__2___closed__1; static lean_object* l_Lean_initFn____x40_Lean_Class___hyg_78____closed__3; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Class___hyg_78____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getOutParamPositions_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isClass___boxed(lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); @@ -3499,7 +3499,7 @@ x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); x_13 = 0; -x_14 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_13); +x_14 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_13); if (x_14 == 0) { lean_object* x_15; lean_object* x_16; uint8_t x_17; diff --git a/stage0/stdlib/Lean/Compiler/CSimpAttr.c b/stage0/stdlib/Lean/Compiler/CSimpAttr.c index 9beec8e7e6f0..9c3e8b2aa653 100644 --- a/stage0/stdlib/Lean/Compiler/CSimpAttr.c +++ b/stage0/stdlib/Lean/Compiler/CSimpAttr.c @@ -119,7 +119,6 @@ lean_object* l_Lean_Expr_constLevels_x21(lean_object*); lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_464____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_ReplaceImpl_Cache_new(lean_object*); -lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__4(lean_object*); lean_object* l_Lean_Expr_app___override(lean_object*, lean_object*); static lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_CSimpAttr_0__Lean_Compiler_CSimp_isConstantReplacement_x3f___spec__1___closed__3; LEAN_EXPORT lean_object* l_Lean_getConstInfo___at___private_Lean_Compiler_CSimpAttr_0__Lean_Compiler_CSimp_isConstantReplacement_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -176,6 +175,7 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_add(lean_object*, uint8_t, lean_o uint8_t l_Lean_Expr_isConst(lean_object*); lean_object* l_Lean_Expr_lam___override(lean_object*, lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____lambda__1(lean_object*, lean_object*); +lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_464____closed__13; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_CSimp_initFn____x40_Lean_Compiler_CSimpAttr___hyg_132____spec__3___closed__2; @@ -370,7 +370,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_6417____spec__4(x_4); +x_6 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(x_4); lean_ctor_set(x_1, 1, x_6); lean_ctor_set(x_1, 0, x_5); return x_1; @@ -384,7 +384,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_6417____spec__4(x_8); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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/IR/ResetReuse.c b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c index cdf9916f30a9..93c6c977634c 100644 --- a/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c +++ b/stage0/stdlib/Lean/Compiler/IR/ResetReuse.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___boxed(lean_object*, lean_object*); @@ -21,6 +22,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_Rese uint8_t lean_usize_dec_le(size_t, size_t); uint8_t l_Lean_IR_HasIndex_visitFnBody(lean_object*, lean_object*); uint64_t lean_uint64_of_nat(lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5; size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); @@ -30,36 +32,37 @@ uint8_t lean_usize_dec_eq(size_t, size_t); lean_object* lean_array_fget(lean_object*, lean_object*); uint8_t l_Lean_IR_CtorInfo_isScalar(lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(lean_object*, lean_object*, uint8_t, lean_object*); uint8_t l_Lean_IR_FnBody_isTerminal(lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_FnBody_hasLiveVar(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuseCore(lean_object*, uint8_t); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_IR_ResetReuse_R___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1___boxed(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_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(lean_object*, uint8_t, size_t, size_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_Context_lctx___default; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_R(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_getPrefix(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_FnBody_body(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_collectResets(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dfinalize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_IR_ResetReuse_R___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_IR_ResetReuse_R___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1; lean_object* l_Lean_IR_instBEqVarId___boxed(lean_object*, lean_object*); @@ -67,519 +70,559 @@ static size_t l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___s LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dfinalize___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___spec__1(lean_object*, lean_object*, size_t, size_t); lean_object* lean_usize_to_nat(size_t); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1(lean_object*, lean_object*, uint8_t, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object*); uint8_t l_Lean_IR_FnBody_beq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_instHashableVarId___boxed(lean_object*); -static lean_object* l_Lean_IR_Decl_insertResetReuse___closed__1; +LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuseCore___boxed(lean_object*, lean_object*); lean_object* l_Lean_IR_FnBody_setBody(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_IR_Decl_updateBody_x21(lean_object*, lean_object*); -static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3; LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_D(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_Context_alreadyFound___default; +lean_object* l_Lean_IR_AltCore_body(lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_argsContainsVar___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* l_Lean_IR_MaxIndex_collectDecl(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__4; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh(lean_object*); lean_object* lean_array_uget(lean_object*, size_t); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_IR_LocalContext_addJP(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1; +uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(lean_object*, lean_object*, uint8_t, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2; lean_object* lean_nat_add(lean_object*, lean_object*); +static lean_object* l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_Context_casesVars___default; -LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object*, lean_object*, uint8_t); lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +LEAN_EXPORT uint8_t l_Lean_IR_ResetReuse_Context_relaxedReuse___default; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT uint8_t l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(lean_object* x_1, lean_object* x_2, uint8_t x_3) { _start: { -lean_object* x_3; lean_object* x_4; uint8_t x_5; -x_3 = lean_ctor_get(x_1, 2); -x_4 = lean_ctor_get(x_2, 2); -x_5 = lean_nat_dec_eq(x_3, x_4); -if (x_5 == 0) +lean_object* x_4; lean_object* x_5; uint8_t x_6; +x_4 = lean_ctor_get(x_1, 2); +x_5 = lean_ctor_get(x_2, 2); +x_6 = lean_nat_dec_eq(x_4, x_5); +if (x_6 == 0) { -uint8_t x_6; -x_6 = 0; -return x_6; +uint8_t x_7; +x_7 = 0; +return x_7; } else { -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = lean_ctor_get(x_1, 3); -x_8 = lean_ctor_get(x_2, 3); -x_9 = lean_nat_dec_eq(x_7, x_8); -if (x_9 == 0) +lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_8 = lean_ctor_get(x_1, 3); +x_9 = lean_ctor_get(x_2, 3); +x_10 = lean_nat_dec_eq(x_8, x_9); +if (x_10 == 0) { -uint8_t x_10; -x_10 = 0; -return x_10; +uint8_t x_11; +x_11 = 0; +return x_11; } else { -lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_11 = lean_ctor_get(x_1, 4); -x_12 = lean_ctor_get(x_2, 4); -x_13 = lean_nat_dec_eq(x_11, x_12); -if (x_13 == 0) +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_1, 4); +x_13 = lean_ctor_get(x_2, 4); +x_14 = lean_nat_dec_eq(x_12, x_13); +if (x_14 == 0) { -uint8_t x_14; -x_14 = 0; -return x_14; +uint8_t x_15; +x_15 = 0; +return x_15; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; uint8_t x_19; -x_15 = lean_ctor_get(x_1, 0); -x_16 = l_Lean_Name_getPrefix(x_15); -x_17 = lean_ctor_get(x_2, 0); -x_18 = l_Lean_Name_getPrefix(x_17); -x_19 = lean_name_eq(x_16, x_18); -lean_dec(x_18); -lean_dec(x_16); -return x_19; +if (x_3 == 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_1, 0); +x_17 = l_Lean_Name_getPrefix(x_16); +x_18 = lean_ctor_get(x_2, 0); +x_19 = l_Lean_Name_getPrefix(x_18); +x_20 = lean_name_eq(x_17, x_19); +lean_dec(x_19); +lean_dec(x_17); +return x_20; +} +else +{ +uint8_t x_21; +x_21 = 1; +return x_21; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse___boxed(lean_object* x_1, lean_object* x_2) { +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_1, x_2); +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_3); +lean_dec(x_3); +x_5 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_1, x_2, x_4); lean_dec(x_2); lean_dec(x_1); -x_4 = lean_box(x_3); -return x_4; +x_6 = lean_box(x_5); +return x_6; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1(lean_object* x_1, lean_object* x_2, uint8_t x_3, size_t x_4, size_t x_5, lean_object* x_6) { _start: { -uint8_t x_6; -x_6 = lean_usize_dec_lt(x_4, x_3); -if (x_6 == 0) +uint8_t x_7; +x_7 = lean_usize_dec_lt(x_5, x_4); +if (x_7 == 0) { lean_dec(x_1); -return x_5; +return x_6; } else { -lean_object* x_7; lean_object* x_8; lean_object* x_9; size_t x_10; size_t x_11; -x_7 = lean_array_uget(x_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); -if (lean_obj_tag(x_7) == 0) +lean_object* x_8; lean_object* x_9; lean_object* x_10; size_t x_11; size_t x_12; +x_8 = lean_array_uget(x_6, x_5); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_array_uset(x_6, x_5, x_9); +x_11 = 1; +x_12 = lean_usize_add(x_5, x_11); +if (lean_obj_tag(x_8) == 0) { -uint8_t x_12; -x_12 = !lean_is_exclusive(x_7); -if (x_12 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_8); +if (x_13 == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; -x_13 = lean_ctor_get(x_7, 1); +lean_object* x_14; lean_object* x_15; lean_object* x_16; +x_14 = lean_ctor_get(x_8, 1); lean_inc(x_1); -x_14 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_13); -lean_ctor_set(x_7, 1, x_14); -x_15 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_15; +x_15 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_14); +lean_ctor_set(x_8, 1, x_15); +x_16 = lean_array_uset(x_10, x_5, x_8); +x_5 = x_12; +x_6 = x_16; goto _start; } else { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_17 = lean_ctor_get(x_7, 0); -x_18 = lean_ctor_get(x_7, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +lean_inc(x_19); lean_inc(x_18); -lean_inc(x_17); -lean_dec(x_7); +lean_dec(x_8); lean_inc(x_1); -x_19 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_17); -lean_ctor_set(x_20, 1, x_19); -x_21 = lean_array_uset(x_9, x_4, x_20); -x_4 = x_11; -x_5 = x_21; +x_20 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_19); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_18); +lean_ctor_set(x_21, 1, x_20); +x_22 = lean_array_uset(x_10, x_5, x_21); +x_5 = x_12; +x_6 = x_22; goto _start; } } else { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_7); -if (x_23 == 0) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_8); +if (x_24 == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_7, 0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_8, 0); lean_inc(x_1); -x_25 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_24); -lean_ctor_set(x_7, 0, x_25); -x_26 = lean_array_uset(x_9, x_4, x_7); -x_4 = x_11; -x_5 = x_26; +x_26 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_25); +lean_ctor_set(x_8, 0, x_26); +x_27 = lean_array_uset(x_10, x_5, x_8); +x_5 = x_12; +x_6 = x_27; goto _start; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_7, 0); -lean_inc(x_28); -lean_dec(x_7); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_8, 0); +lean_inc(x_29); +lean_dec(x_8); lean_inc(x_1); -x_29 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_28); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_29); -x_31 = lean_array_uset(x_9, x_4, x_30); -x_4 = x_11; -x_5 = x_31; +x_30 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_29); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_30); +x_32 = lean_array_uset(x_10, x_5, x_31); +x_5 = x_12; +x_6 = x_32; goto _start; } } } } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4) { _start: { -switch (lean_obj_tag(x_3)) { +switch (lean_obj_tag(x_4)) { case 0: { -lean_object* x_4; -x_4 = lean_ctor_get(x_3, 2); -lean_inc(x_4); -if (lean_obj_tag(x_4) == 0) +lean_object* x_5; +x_5 = lean_ctor_get(x_4, 2); +lean_inc(x_5); +if (lean_obj_tag(x_5) == 0) { -uint8_t x_5; -x_5 = !lean_is_exclusive(x_3); -if (x_5 == 0) +uint8_t x_6; +x_6 = !lean_is_exclusive(x_4); +if (x_6 == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_6 = lean_ctor_get(x_3, 3); -x_7 = lean_ctor_get(x_3, 2); -lean_dec(x_7); -x_8 = lean_ctor_get(x_4, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_4, 1); +lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; +x_7 = lean_ctor_get(x_4, 3); +x_8 = lean_ctor_get(x_4, 2); +lean_dec(x_8); +x_9 = lean_ctor_get(x_5, 0); lean_inc(x_9); -x_10 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_8); -if (x_10 == 0) +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +x_11 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_9, x_3); +if (x_11 == 0) { -lean_object* x_11; +lean_object* x_12; +lean_dec(x_10); lean_dec(x_9); -lean_dec(x_8); -x_11 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_6); -lean_ctor_set(x_3, 3, x_11); -return x_3; +x_12 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_7); +lean_ctor_set(x_4, 3, x_12); +return x_4; } else { -lean_object* x_12; lean_object* x_13; uint8_t x_14; -lean_dec(x_4); -x_12 = lean_ctor_get(x_2, 1); -x_13 = lean_ctor_get(x_8, 1); -lean_inc(x_13); -x_14 = lean_nat_dec_eq(x_12, x_13); -lean_dec(x_13); -if (x_14 == 0) -{ -uint8_t x_15; lean_object* x_16; -x_15 = 1; -x_16 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_16, 1, x_8); -lean_ctor_set(x_16, 2, x_9); -lean_ctor_set_uint8(x_16, sizeof(void*)*3, x_15); -lean_ctor_set(x_3, 2, x_16); -return x_3; +lean_object* x_13; lean_object* x_14; uint8_t x_15; +lean_dec(x_5); +x_13 = lean_ctor_get(x_2, 1); +x_14 = lean_ctor_get(x_9, 1); +lean_inc(x_14); +x_15 = lean_nat_dec_eq(x_13, x_14); +lean_dec(x_14); +if (x_15 == 0) +{ +uint8_t x_16; lean_object* x_17; +x_16 = 1; +x_17 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_17, 0, x_1); +lean_ctor_set(x_17, 1, x_9); +lean_ctor_set(x_17, 2, x_10); +lean_ctor_set_uint8(x_17, sizeof(void*)*3, x_16); +lean_ctor_set(x_4, 2, x_17); +return x_4; } else { -uint8_t x_17; lean_object* x_18; -x_17 = 0; -x_18 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_18, 0, x_1); -lean_ctor_set(x_18, 1, x_8); -lean_ctor_set(x_18, 2, x_9); -lean_ctor_set_uint8(x_18, sizeof(void*)*3, x_17); -lean_ctor_set(x_3, 2, x_18); -return x_3; +uint8_t x_18; lean_object* x_19; +x_18 = 0; +x_19 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_19, 0, x_1); +lean_ctor_set(x_19, 1, x_9); +lean_ctor_set(x_19, 2, x_10); +lean_ctor_set_uint8(x_19, sizeof(void*)*3, x_18); +lean_ctor_set(x_4, 2, x_19); +return x_4; } } } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_19 = lean_ctor_get(x_3, 0); -x_20 = lean_ctor_get(x_3, 1); -x_21 = lean_ctor_get(x_3, 3); +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_20 = lean_ctor_get(x_4, 0); +x_21 = lean_ctor_get(x_4, 1); +x_22 = lean_ctor_get(x_4, 3); +lean_inc(x_22); lean_inc(x_21); lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_3); -x_22 = lean_ctor_get(x_4, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_4, 1); +lean_dec(x_4); +x_23 = lean_ctor_get(x_5, 0); lean_inc(x_23); -x_24 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_22); -if (x_24 == 0) +x_24 = lean_ctor_get(x_5, 1); +lean_inc(x_24); +x_25 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mayReuse(x_2, x_23, x_3); +if (x_25 == 0) { -lean_object* x_25; lean_object* x_26; +lean_object* x_26; lean_object* x_27; +lean_dec(x_24); lean_dec(x_23); -lean_dec(x_22); -x_25 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_21); -x_26 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_26, 0, x_19); -lean_ctor_set(x_26, 1, x_20); -lean_ctor_set(x_26, 2, x_4); -lean_ctor_set(x_26, 3, x_25); -return x_26; +x_26 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_22); +x_27 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_27, 0, x_20); +lean_ctor_set(x_27, 1, x_21); +lean_ctor_set(x_27, 2, x_5); +lean_ctor_set(x_27, 3, x_26); +return x_27; } else { -lean_object* x_27; lean_object* x_28; uint8_t x_29; -lean_dec(x_4); -x_27 = lean_ctor_get(x_2, 1); -x_28 = lean_ctor_get(x_22, 1); -lean_inc(x_28); -x_29 = lean_nat_dec_eq(x_27, x_28); -lean_dec(x_28); -if (x_29 == 0) -{ -uint8_t x_30; lean_object* x_31; lean_object* x_32; -x_30 = 1; -x_31 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_31, 0, x_1); -lean_ctor_set(x_31, 1, x_22); -lean_ctor_set(x_31, 2, x_23); -lean_ctor_set_uint8(x_31, sizeof(void*)*3, x_30); -x_32 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_32, 0, x_19); -lean_ctor_set(x_32, 1, x_20); -lean_ctor_set(x_32, 2, x_31); -lean_ctor_set(x_32, 3, x_21); -return x_32; +lean_object* x_28; lean_object* x_29; uint8_t x_30; +lean_dec(x_5); +x_28 = lean_ctor_get(x_2, 1); +x_29 = lean_ctor_get(x_23, 1); +lean_inc(x_29); +x_30 = lean_nat_dec_eq(x_28, x_29); +lean_dec(x_29); +if (x_30 == 0) +{ +uint8_t x_31; lean_object* x_32; lean_object* x_33; +x_31 = 1; +x_32 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_32, 0, x_1); +lean_ctor_set(x_32, 1, x_23); +lean_ctor_set(x_32, 2, x_24); +lean_ctor_set_uint8(x_32, sizeof(void*)*3, x_31); +x_33 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_33, 0, x_20); +lean_ctor_set(x_33, 1, x_21); +lean_ctor_set(x_33, 2, x_32); +lean_ctor_set(x_33, 3, x_22); +return x_33; } else { -uint8_t x_33; lean_object* x_34; lean_object* x_35; -x_33 = 0; -x_34 = lean_alloc_ctor(2, 3, 1); -lean_ctor_set(x_34, 0, x_1); -lean_ctor_set(x_34, 1, x_22); -lean_ctor_set(x_34, 2, x_23); -lean_ctor_set_uint8(x_34, sizeof(void*)*3, x_33); -x_35 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_35, 0, x_19); -lean_ctor_set(x_35, 1, x_20); -lean_ctor_set(x_35, 2, x_34); -lean_ctor_set(x_35, 3, x_21); -return x_35; +uint8_t x_34; lean_object* x_35; lean_object* x_36; +x_34 = 0; +x_35 = lean_alloc_ctor(2, 3, 1); +lean_ctor_set(x_35, 0, x_1); +lean_ctor_set(x_35, 1, x_23); +lean_ctor_set(x_35, 2, x_24); +lean_ctor_set_uint8(x_35, sizeof(void*)*3, x_34); +x_36 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_36, 0, x_20); +lean_ctor_set(x_36, 1, x_21); +lean_ctor_set(x_36, 2, x_35); +lean_ctor_set(x_36, 3, x_22); +return x_36; } } } } else { -uint8_t x_36; -lean_dec(x_4); -x_36 = l_Lean_IR_FnBody_isTerminal(x_3); -if (x_36 == 0) +uint8_t x_37; +lean_dec(x_5); +x_37 = l_Lean_IR_FnBody_isTerminal(x_4); +if (x_37 == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -x_37 = l_Lean_IR_FnBody_body(x_3); -x_38 = lean_box(13); -x_39 = l_Lean_IR_FnBody_setBody(x_3, x_38); -x_40 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_37); -x_41 = l_Lean_IR_FnBody_setBody(x_39, x_40); -return x_41; +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_38 = l_Lean_IR_FnBody_body(x_4); +x_39 = lean_box(13); +x_40 = l_Lean_IR_FnBody_setBody(x_4, x_39); +x_41 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_38); +x_42 = l_Lean_IR_FnBody_setBody(x_40, x_41); +return x_42; } else { lean_dec(x_1); -return x_3; +return x_4; } } } case 1: { -uint8_t x_42; -x_42 = !lean_is_exclusive(x_3); -if (x_42 == 0) +uint8_t x_43; +x_43 = !lean_is_exclusive(x_4); +if (x_43 == 0) { -lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; -x_43 = lean_ctor_get(x_3, 2); -x_44 = lean_ctor_get(x_3, 3); -lean_inc(x_43); +lean_object* x_44; lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_44 = lean_ctor_get(x_4, 2); +x_45 = lean_ctor_get(x_4, 3); +lean_inc(x_44); lean_inc(x_1); -x_45 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_43); -lean_inc(x_45); -lean_inc(x_43); -x_46 = l_Lean_IR_FnBody_beq(x_43, x_45); -if (x_46 == 0) +x_46 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_44); +lean_inc(x_46); +lean_inc(x_44); +x_47 = l_Lean_IR_FnBody_beq(x_44, x_46); +if (x_47 == 0) { -lean_dec(x_43); +lean_dec(x_44); lean_dec(x_1); -lean_ctor_set(x_3, 2, x_45); -return x_3; +lean_ctor_set(x_4, 2, x_46); +return x_4; } else { -lean_object* x_47; -lean_dec(x_45); -x_47 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_44); -lean_ctor_set(x_3, 3, x_47); -return x_3; +lean_object* x_48; +lean_dec(x_46); +x_48 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_45); +lean_ctor_set(x_4, 3, x_48); +return x_4; } } else { -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_3, 0); -x_49 = lean_ctor_get(x_3, 1); -x_50 = lean_ctor_get(x_3, 2); -x_51 = lean_ctor_get(x_3, 3); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; +x_49 = lean_ctor_get(x_4, 0); +x_50 = lean_ctor_get(x_4, 1); +x_51 = lean_ctor_get(x_4, 2); +x_52 = lean_ctor_get(x_4, 3); +lean_inc(x_52); lean_inc(x_51); lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_3); -lean_inc(x_50); +lean_dec(x_4); +lean_inc(x_51); lean_inc(x_1); -x_52 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_50); -lean_inc(x_52); -lean_inc(x_50); -x_53 = l_Lean_IR_FnBody_beq(x_50, x_52); -if (x_53 == 0) +x_53 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_51); +lean_inc(x_53); +lean_inc(x_51); +x_54 = l_Lean_IR_FnBody_beq(x_51, x_53); +if (x_54 == 0) { -lean_object* x_54; -lean_dec(x_50); +lean_object* x_55; +lean_dec(x_51); lean_dec(x_1); -x_54 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_54, 0, x_48); -lean_ctor_set(x_54, 1, x_49); -lean_ctor_set(x_54, 2, x_52); -lean_ctor_set(x_54, 3, x_51); -return x_54; +x_55 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_55, 0, x_49); +lean_ctor_set(x_55, 1, x_50); +lean_ctor_set(x_55, 2, x_53); +lean_ctor_set(x_55, 3, x_52); +return x_55; } else { -lean_object* x_55; lean_object* x_56; -lean_dec(x_52); -x_55 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_51); -x_56 = lean_alloc_ctor(1, 4, 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); -return x_56; +lean_object* x_56; lean_object* x_57; +lean_dec(x_53); +x_56 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_52); +x_57 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_57, 0, x_49); +lean_ctor_set(x_57, 1, x_50); +lean_ctor_set(x_57, 2, x_51); +lean_ctor_set(x_57, 3, x_56); +return x_57; } } } case 10: { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_3); -if (x_57 == 0) +uint8_t x_58; +x_58 = !lean_is_exclusive(x_4); +if (x_58 == 0) { -lean_object* x_58; lean_object* x_59; size_t x_60; size_t x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_3, 3); -x_59 = lean_array_get_size(x_58); -x_60 = lean_usize_of_nat(x_59); -lean_dec(x_59); -x_61 = 0; -x_62 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(x_1, x_2, x_60, x_61, x_58); -lean_ctor_set(x_3, 3, x_62); -return x_3; +lean_object* x_59; lean_object* x_60; size_t x_61; size_t x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_4, 3); +x_60 = lean_array_get_size(x_59); +x_61 = lean_usize_of_nat(x_60); +lean_dec(x_60); +x_62 = 0; +x_63 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1(x_1, x_2, x_3, x_61, x_62, x_59); +lean_ctor_set(x_4, 3, x_63); +return x_4; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; size_t x_68; size_t x_69; lean_object* x_70; lean_object* x_71; -x_63 = lean_ctor_get(x_3, 0); -x_64 = lean_ctor_get(x_3, 1); -x_65 = lean_ctor_get(x_3, 2); -x_66 = lean_ctor_get(x_3, 3); +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; size_t x_69; size_t x_70; lean_object* x_71; lean_object* x_72; +x_64 = lean_ctor_get(x_4, 0); +x_65 = lean_ctor_get(x_4, 1); +x_66 = lean_ctor_get(x_4, 2); +x_67 = lean_ctor_get(x_4, 3); +lean_inc(x_67); lean_inc(x_66); lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_3); -x_67 = lean_array_get_size(x_66); -x_68 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_69 = 0; -x_70 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(x_1, x_2, x_68, x_69, x_66); -x_71 = lean_alloc_ctor(10, 4, 0); -lean_ctor_set(x_71, 0, x_63); -lean_ctor_set(x_71, 1, x_64); -lean_ctor_set(x_71, 2, x_65); -lean_ctor_set(x_71, 3, x_70); -return x_71; +lean_dec(x_4); +x_68 = lean_array_get_size(x_67); +x_69 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_70 = 0; +x_71 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1(x_1, x_2, x_3, x_69, x_70, x_67); +x_72 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_72, 0, x_64); +lean_ctor_set(x_72, 1, x_65); +lean_ctor_set(x_72, 2, x_66); +lean_ctor_set(x_72, 3, x_71); +return x_72; } } default: { -uint8_t x_72; -x_72 = l_Lean_IR_FnBody_isTerminal(x_3); -if (x_72 == 0) +uint8_t x_73; +x_73 = l_Lean_IR_FnBody_isTerminal(x_4); +if (x_73 == 0) { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_73 = l_Lean_IR_FnBody_body(x_3); -x_74 = lean_box(13); -x_75 = l_Lean_IR_FnBody_setBody(x_3, x_74); -x_76 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_73); -x_77 = l_Lean_IR_FnBody_setBody(x_75, x_76); -return x_77; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_74 = l_Lean_IR_FnBody_body(x_4); +x_75 = lean_box(13); +x_76 = l_Lean_IR_FnBody_setBody(x_4, x_75); +x_77 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_74); +x_78 = l_Lean_IR_FnBody_setBody(x_76, x_77); +return x_78; } else { lean_dec(x_1); -return x_3; +return x_4; } } } } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__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_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___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: { -size_t x_6; size_t x_7; lean_object* x_8; -x_6 = lean_unbox_usize(x_3); +uint8_t x_7; size_t x_8; size_t x_9; lean_object* x_10; +x_7 = lean_unbox(x_3); lean_dec(x_3); -x_7 = lean_unbox_usize(x_4); +x_8 = lean_unbox_usize(x_4); lean_dec(x_4); -x_8 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___spec__1(x_1, x_2, x_6, x_7, x_5); +x_9 = lean_unbox_usize(x_5); +lean_dec(x_5); +x_10 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___spec__1(x_1, x_2, x_7, x_8, x_9, x_6); lean_dec(x_2); -return x_8; +return x_10; } } -LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_4; -x_4 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_3); +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_3); +lean_dec(x_3); +x_6 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_5, x_4); lean_dec(x_2); -return x_4; +return x_6; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_3); +lean_dec(x_3); +x_6 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_1, x_2, x_5, x_4); +lean_dec(x_2); +return x_6; } } static lean_object* _init_l_Lean_IR_ResetReuse_Context_lctx___default() { @@ -590,7 +633,7 @@ x_1 = lean_box(0); return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1() { _start: { lean_object* x_1; @@ -598,17 +641,17 @@ x_1 = l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_box(0), lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__3() { _start: { lean_object* x_1; @@ -616,7 +659,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_IR_instBEqVarId___boxed), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__4() { _start: { lean_object* x_1; @@ -624,11 +667,11 @@ x_1 = lean_alloc_closure((void*)(l_Lean_IR_instHashableVarId___boxed), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2; x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -636,19 +679,27 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1() { +static lean_object* _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5; return x_1; } } -static lean_object* _init_l_Lean_IR_ResetReuse_Context_casesVars___default() { +static lean_object* _init_l_Lean_IR_ResetReuse_Context_alreadyFound___default() { _start: { lean_object* x_1; -x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; +x_1 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1; +return x_1; +} +} +static uint8_t _init_l_Lean_IR_ResetReuse_Context_relaxedReuse___default() { +_start: +{ +uint8_t x_1; +x_1 = 0; return x_1; } } @@ -689,35 +740,36 @@ x_6 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_mkFresh___ra x_7 = !lean_is_exclusive(x_6); if (x_7 == 0) { -lean_object* x_8; lean_object* x_9; uint8_t x_10; +lean_object* x_8; uint8_t x_9; lean_object* x_10; uint8_t x_11; x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get_uint8(x_4, sizeof(void*)*2); lean_inc(x_3); lean_inc(x_8); -x_9 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_8, x_2, x_3); -lean_inc(x_9); +x_10 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_8, x_2, x_9, x_3); +lean_inc(x_10); lean_inc(x_3); -x_10 = l_Lean_IR_FnBody_beq(x_3, x_9); -if (x_10 == 0) +x_11 = l_Lean_IR_FnBody_beq(x_3, x_10); +if (x_11 == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_dec(x_3); -x_11 = lean_ctor_get(x_2, 2); -lean_inc(x_11); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_1); -x_13 = lean_box(7); -x_14 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_14, 0, x_8); -lean_ctor_set(x_14, 1, x_13); -lean_ctor_set(x_14, 2, x_12); -lean_ctor_set(x_14, 3, x_9); -lean_ctor_set(x_6, 0, x_14); +x_12 = lean_ctor_get(x_2, 2); +lean_inc(x_12); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_1); +x_14 = lean_box(7); +x_15 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_15, 0, x_8); +lean_ctor_set(x_15, 1, x_14); +lean_ctor_set(x_15, 2, x_13); +lean_ctor_set(x_15, 3, x_10); +lean_ctor_set(x_6, 0, x_15); return x_6; } else { -lean_dec(x_9); +lean_dec(x_10); lean_dec(x_8); lean_dec(x_1); lean_ctor_set(x_6, 0, x_3); @@ -726,48 +778,49 @@ return x_6; } else { -lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = lean_ctor_get(x_6, 0); -x_16 = lean_ctor_get(x_6, 1); +lean_object* x_16; lean_object* x_17; uint8_t x_18; lean_object* x_19; uint8_t x_20; +x_16 = lean_ctor_get(x_6, 0); +x_17 = lean_ctor_get(x_6, 1); +lean_inc(x_17); lean_inc(x_16); -lean_inc(x_15); lean_dec(x_6); +x_18 = lean_ctor_get_uint8(x_4, sizeof(void*)*2); lean_inc(x_3); -lean_inc(x_15); -x_17 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S(x_15, x_2, x_3); -lean_inc(x_17); +lean_inc(x_16); +x_19 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_S_go(x_16, x_2, x_18, x_3); +lean_inc(x_19); lean_inc(x_3); -x_18 = l_Lean_IR_FnBody_beq(x_3, x_17); -if (x_18 == 0) +x_20 = l_Lean_IR_FnBody_beq(x_3, x_19); +if (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_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_dec(x_3); -x_19 = lean_ctor_get(x_2, 2); -lean_inc(x_19); -x_20 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_1); -x_21 = lean_box(7); -x_22 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_22, 0, x_15); -lean_ctor_set(x_22, 1, x_21); -lean_ctor_set(x_22, 2, x_20); -lean_ctor_set(x_22, 3, x_17); -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 -{ -lean_object* x_24; -lean_dec(x_17); -lean_dec(x_15); +x_21 = lean_ctor_get(x_2, 2); +lean_inc(x_21); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_1); +x_23 = lean_box(7); +x_24 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_24, 0, x_16); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set(x_24, 2, x_22); +lean_ctor_set(x_24, 3, x_19); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_17); +return x_25; +} +else +{ +lean_object* x_26; +lean_dec(x_19); +lean_dec(x_16); lean_dec(x_1); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_3); -lean_ctor_set(x_24, 1, x_16); -return x_24; +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_3); +lean_ctor_set(x_26, 1, x_17); +return x_26; } } } @@ -1141,7 +1194,7 @@ uint8_t x_6; x_6 = !lean_is_exclusive(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_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; +lean_object* x_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; lean_object* x_20; lean_object* x_21; uint8_t x_22; x_7 = lean_ctor_get(x_3, 0); x_8 = lean_ctor_get(x_3, 1); x_9 = lean_ctor_get(x_3, 2); @@ -1150,635 +1203,639 @@ x_11 = lean_ctor_get(x_4, 0); lean_inc(x_11); x_12 = lean_ctor_get(x_4, 1); lean_inc(x_12); +x_13 = lean_ctor_get_uint8(x_4, sizeof(void*)*2); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_13 = l_Lean_IR_LocalContext_addJP(x_11, x_7, x_8, x_9); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_12); +x_14 = l_Lean_IR_LocalContext_addJP(x_11, x_7, x_8, x_9); +x_15 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_15, 0, x_14); +lean_ctor_set(x_15, 1, x_12); +lean_ctor_set_uint8(x_15, sizeof(void*)*2, x_13); lean_inc(x_1); -x_15 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_10, x_14, x_5); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +x_16 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_10, x_15, x_5); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_ctor_get(x_16, 0); +x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); lean_dec(x_16); -x_20 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_9, x_4, x_17); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +x_19 = lean_ctor_get(x_17, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_17, 1); +lean_inc(x_20); +lean_dec(x_17); +x_21 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_9, x_4, x_18); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) { -lean_object* x_22; uint8_t x_23; -x_22 = lean_ctor_get(x_20, 0); -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); -x_25 = lean_ctor_get(x_22, 1); -lean_dec(x_25); -lean_ctor_set(x_3, 3, x_18); -lean_ctor_set(x_3, 2, x_24); -lean_ctor_set(x_22, 1, x_19); -lean_ctor_set(x_22, 0, x_3); -return x_20; +lean_object* x_23; uint8_t x_24; +x_23 = lean_ctor_get(x_21, 0); +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); +x_26 = lean_ctor_get(x_23, 1); +lean_dec(x_26); +lean_ctor_set(x_3, 3, x_19); +lean_ctor_set(x_3, 2, x_25); +lean_ctor_set(x_23, 1, x_20); +lean_ctor_set(x_23, 0, x_3); +return x_21; } else { -lean_object* x_26; lean_object* x_27; -x_26 = lean_ctor_get(x_22, 0); -lean_inc(x_26); -lean_dec(x_22); -lean_ctor_set(x_3, 3, x_18); -lean_ctor_set(x_3, 2, x_26); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_3); -lean_ctor_set(x_27, 1, x_19); -lean_ctor_set(x_20, 0, x_27); -return x_20; +lean_object* x_27; lean_object* x_28; +x_27 = lean_ctor_get(x_23, 0); +lean_inc(x_27); +lean_dec(x_23); +lean_ctor_set(x_3, 3, x_19); +lean_ctor_set(x_3, 2, x_27); +x_28 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_28, 0, x_3); +lean_ctor_set(x_28, 1, x_20); +lean_ctor_set(x_21, 0, x_28); +return x_21; } } 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; -x_28 = lean_ctor_get(x_20, 0); -x_29 = lean_ctor_get(x_20, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_20); -x_30 = lean_ctor_get(x_28, 0); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); lean_inc(x_30); -if (lean_is_exclusive(x_28)) { - lean_ctor_release(x_28, 0); - lean_ctor_release(x_28, 1); - x_31 = x_28; +lean_inc(x_29); +lean_dec(x_21); +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_32 = x_29; } else { - lean_dec_ref(x_28); - x_31 = lean_box(0); + lean_dec_ref(x_29); + x_32 = lean_box(0); } -lean_ctor_set(x_3, 3, x_18); -lean_ctor_set(x_3, 2, x_30); -if (lean_is_scalar(x_31)) { - x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_3, 3, x_19); +lean_ctor_set(x_3, 2, x_31); +if (lean_is_scalar(x_32)) { + x_33 = lean_alloc_ctor(0, 2, 0); } else { - x_32 = x_31; + x_33 = x_32; } -lean_ctor_set(x_32, 0, x_3); -lean_ctor_set(x_32, 1, x_19); -x_33 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_33, 0, x_32); -lean_ctor_set(x_33, 1, x_29); -return x_33; +lean_ctor_set(x_33, 0, x_3); +lean_ctor_set(x_33, 1, x_20); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_30); +return x_34; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_34 = lean_ctor_get(x_3, 0); -x_35 = lean_ctor_get(x_3, 1); -x_36 = lean_ctor_get(x_3, 2); -x_37 = lean_ctor_get(x_3, 3); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_35 = lean_ctor_get(x_3, 0); +x_36 = lean_ctor_get(x_3, 1); +x_37 = lean_ctor_get(x_3, 2); +x_38 = lean_ctor_get(x_3, 3); +lean_inc(x_38); lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); lean_dec(x_3); -x_38 = lean_ctor_get(x_4, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_4, 1); +x_39 = lean_ctor_get(x_4, 0); lean_inc(x_39); +x_40 = lean_ctor_get(x_4, 1); +lean_inc(x_40); +x_41 = lean_ctor_get_uint8(x_4, sizeof(void*)*2); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -x_40 = l_Lean_IR_LocalContext_addJP(x_38, x_34, x_35, x_36); -x_41 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_39); +x_42 = l_Lean_IR_LocalContext_addJP(x_39, x_35, x_36, x_37); +x_43 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_40); +lean_ctor_set_uint8(x_43, sizeof(void*)*2, x_41); lean_inc(x_1); -x_42 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_37, x_41, x_5); -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); +x_44 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_38, x_43, x_5); +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -x_46 = lean_ctor_get(x_43, 1); +x_46 = lean_ctor_get(x_44, 1); lean_inc(x_46); -lean_dec(x_43); -x_47 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_36, x_4, x_44); -x_48 = lean_ctor_get(x_47, 0); +lean_dec(x_44); +x_47 = lean_ctor_get(x_45, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_45, 1); lean_inc(x_48); -x_49 = lean_ctor_get(x_47, 1); -lean_inc(x_49); -if (lean_is_exclusive(x_47)) { - lean_ctor_release(x_47, 0); - lean_ctor_release(x_47, 1); - x_50 = x_47; -} else { - lean_dec_ref(x_47); - x_50 = lean_box(0); -} -x_51 = lean_ctor_get(x_48, 0); +lean_dec(x_45); +x_49 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_37, x_4, x_46); +x_50 = lean_ctor_get(x_49, 0); +lean_inc(x_50); +x_51 = lean_ctor_get(x_49, 1); lean_inc(x_51); -if (lean_is_exclusive(x_48)) { - lean_ctor_release(x_48, 0); - lean_ctor_release(x_48, 1); - x_52 = x_48; +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_52 = x_49; } else { - lean_dec_ref(x_48); + lean_dec_ref(x_49); x_52 = lean_box(0); } -x_53 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_53, 0, x_34); -lean_ctor_set(x_53, 1, x_35); -lean_ctor_set(x_53, 2, x_51); -lean_ctor_set(x_53, 3, x_45); -if (lean_is_scalar(x_52)) { - x_54 = lean_alloc_ctor(0, 2, 0); +x_53 = lean_ctor_get(x_50, 0); +lean_inc(x_53); +if (lean_is_exclusive(x_50)) { + lean_ctor_release(x_50, 0); + lean_ctor_release(x_50, 1); + x_54 = x_50; } else { - x_54 = x_52; + lean_dec_ref(x_50); + x_54 = lean_box(0); +} +x_55 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_55, 0, x_35); +lean_ctor_set(x_55, 1, x_36); +lean_ctor_set(x_55, 2, x_53); +lean_ctor_set(x_55, 3, x_47); +if (lean_is_scalar(x_54)) { + x_56 = lean_alloc_ctor(0, 2, 0); +} else { + x_56 = x_54; } -lean_ctor_set(x_54, 0, x_53); -lean_ctor_set(x_54, 1, x_46); -if (lean_is_scalar(x_50)) { - x_55 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_48); +if (lean_is_scalar(x_52)) { + x_57 = lean_alloc_ctor(0, 2, 0); } else { - x_55 = x_50; + x_57 = x_52; } -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_49); -return x_55; +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_51); +return x_57; } } case 10: { -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; -x_56 = lean_ctor_get(x_3, 0); -lean_inc(x_56); -x_57 = lean_ctor_get(x_3, 1); -lean_inc(x_57); -x_58 = lean_ctor_get(x_3, 2); +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; +x_58 = lean_ctor_get(x_3, 0); lean_inc(x_58); -x_59 = lean_ctor_get(x_3, 3); +x_59 = lean_ctor_get(x_3, 1); lean_inc(x_59); -x_60 = lean_ctor_get(x_4, 0); +x_60 = lean_ctor_get(x_3, 2); lean_inc(x_60); +x_61 = lean_ctor_get(x_3, 3); +lean_inc(x_61); +x_62 = lean_ctor_get(x_4, 0); +lean_inc(x_62); lean_inc(x_3); -x_61 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_60, x_1); -x_62 = lean_unbox(x_61); -lean_dec(x_61); -if (x_62 == 0) +x_63 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_62, x_1); +x_64 = lean_unbox(x_63); +lean_dec(x_63); +if (x_64 == 0) { -uint8_t x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +lean_dec(x_61); +lean_dec(x_60); lean_dec(x_59); lean_dec(x_58); -lean_dec(x_57); -lean_dec(x_56); lean_dec(x_4); -lean_dec(x_1); -x_63 = 0; -x_64 = lean_box(x_63); -x_65 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_65, 0, x_3); -lean_ctor_set(x_65, 1, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_5); -return x_66; +lean_dec(x_1); +x_65 = 0; +x_66 = lean_box(x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_3); +lean_ctor_set(x_67, 1, x_66); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_67); +lean_ctor_set(x_68, 1, x_5); +return x_68; } else { -uint8_t x_67; -x_67 = !lean_is_exclusive(x_3); -if (x_67 == 0) +uint8_t x_69; +x_69 = !lean_is_exclusive(x_3); +if (x_69 == 0) { -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; uint8_t x_76; -x_68 = lean_ctor_get(x_3, 3); -lean_dec(x_68); -x_69 = lean_ctor_get(x_3, 2); -lean_dec(x_69); -x_70 = lean_ctor_get(x_3, 1); +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; uint8_t x_78; +x_70 = lean_ctor_get(x_3, 3); lean_dec(x_70); -x_71 = lean_ctor_get(x_3, 0); +x_71 = lean_ctor_get(x_3, 2); lean_dec(x_71); -x_72 = lean_array_get_size(x_59); -x_73 = lean_usize_of_nat(x_72); +x_72 = lean_ctor_get(x_3, 1); lean_dec(x_72); -x_74 = 0; -x_75 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_73, x_74, x_59, x_4, x_5); -x_76 = !lean_is_exclusive(x_75); -if (x_76 == 0) -{ -lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; -x_77 = lean_ctor_get(x_75, 0); -x_78 = lean_ctor_get(x_75, 1); -lean_ctor_set(x_3, 3, x_77); -x_79 = 1; -x_80 = lean_box(x_79); -lean_ctor_set(x_75, 1, x_80); -lean_ctor_set(x_75, 0, x_3); -x_81 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_81, 0, x_75); -lean_ctor_set(x_81, 1, x_78); -return x_81; -} -else -{ -lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; -x_82 = lean_ctor_get(x_75, 0); -x_83 = lean_ctor_get(x_75, 1); -lean_inc(x_83); -lean_inc(x_82); -lean_dec(x_75); -lean_ctor_set(x_3, 3, x_82); -x_84 = 1; -x_85 = lean_box(x_84); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_3); -lean_ctor_set(x_86, 1, x_85); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_83); -return x_87; +x_73 = lean_ctor_get(x_3, 0); +lean_dec(x_73); +x_74 = lean_array_get_size(x_61); +x_75 = lean_usize_of_nat(x_74); +lean_dec(x_74); +x_76 = 0; +x_77 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_75, x_76, x_61, x_4, x_5); +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) +{ +lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; lean_object* x_83; +x_79 = lean_ctor_get(x_77, 0); +x_80 = lean_ctor_get(x_77, 1); +lean_ctor_set(x_3, 3, x_79); +x_81 = 1; +x_82 = lean_box(x_81); +lean_ctor_set(x_77, 1, x_82); +lean_ctor_set(x_77, 0, x_3); +x_83 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_83, 0, x_77); +lean_ctor_set(x_83, 1, x_80); +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 = lean_ctor_get(x_77, 0); +x_85 = lean_ctor_get(x_77, 1); +lean_inc(x_85); +lean_inc(x_84); +lean_dec(x_77); +lean_ctor_set(x_3, 3, x_84); +x_86 = 1; +x_87 = lean_box(x_86); +x_88 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_88, 0, x_3); +lean_ctor_set(x_88, 1, x_87); +x_89 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_89, 0, x_88); +lean_ctor_set(x_89, 1, x_85); +return x_89; } } else { -lean_object* x_88; size_t x_89; size_t 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_90; size_t x_91; size_t x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; uint8_t x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_dec(x_3); -x_88 = lean_array_get_size(x_59); -x_89 = lean_usize_of_nat(x_88); -lean_dec(x_88); -x_90 = 0; -x_91 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_89, x_90, x_59, x_4, x_5); -x_92 = lean_ctor_get(x_91, 0); -lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_94 = x_91; +x_90 = lean_array_get_size(x_61); +x_91 = lean_usize_of_nat(x_90); +lean_dec(x_90); +x_92 = 0; +x_93 = l_Array_mapMUnsafe_map___at___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain___spec__1(x_1, x_2, x_91, x_92, x_61, x_4, x_5); +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; } else { - lean_dec_ref(x_91); - x_94 = lean_box(0); -} -x_95 = lean_alloc_ctor(10, 4, 0); -lean_ctor_set(x_95, 0, x_56); -lean_ctor_set(x_95, 1, x_57); -lean_ctor_set(x_95, 2, x_58); -lean_ctor_set(x_95, 3, x_92); -x_96 = 1; -x_97 = lean_box(x_96); -if (lean_is_scalar(x_94)) { - x_98 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_93); + x_96 = lean_box(0); +} +x_97 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_97, 0, x_58); +lean_ctor_set(x_97, 1, x_59); +lean_ctor_set(x_97, 2, x_60); +lean_ctor_set(x_97, 3, x_94); +x_98 = 1; +x_99 = lean_box(x_98); +if (lean_is_scalar(x_96)) { + x_100 = lean_alloc_ctor(0, 2, 0); } else { - x_98 = x_94; + x_100 = x_96; } -lean_ctor_set(x_98, 0, x_95); -lean_ctor_set(x_98, 1, x_97); -x_99 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_99, 0, x_98); -lean_ctor_set(x_99, 1, x_93); -return x_99; +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_99); +x_101 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_95); +return x_101; } } } default: { -uint8_t x_100; -x_100 = l_Lean_IR_FnBody_isTerminal(x_3); -if (x_100 == 0) +uint8_t x_102; +x_102 = l_Lean_IR_FnBody_isTerminal(x_3); +if (x_102 == 0) { -lean_object* x_101; lean_object* x_102; lean_object* x_103; uint8_t x_104; -x_101 = l_Lean_IR_FnBody_body(x_3); -x_102 = lean_box(13); +lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; +x_103 = l_Lean_IR_FnBody_body(x_3); +x_104 = lean_box(13); lean_inc(x_3); -x_103 = l_Lean_IR_FnBody_setBody(x_3, x_102); -x_104 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(x_103, x_1); -if (x_104 == 0) +x_105 = l_Lean_IR_FnBody_setBody(x_3, x_104); +x_106 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_isCtorUsing(x_105, x_1); +if (x_106 == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; uint8_t x_108; +lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; lean_dec(x_3); lean_inc(x_4); lean_inc(x_1); -x_105 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_101, x_4, x_5); -x_106 = lean_ctor_get(x_105, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_106, 1); -lean_inc(x_107); -x_108 = lean_unbox(x_107); -if (x_108 == 0) -{ -uint8_t x_109; -x_109 = !lean_is_exclusive(x_105); -if (x_109 == 0) -{ -lean_object* x_110; lean_object* x_111; uint8_t x_112; -x_110 = lean_ctor_get(x_105, 1); -x_111 = lean_ctor_get(x_105, 0); -lean_dec(x_111); -x_112 = !lean_is_exclusive(x_106); -if (x_112 == 0) -{ -lean_object* x_113; lean_object* x_114; uint8_t x_115; -x_113 = lean_ctor_get(x_106, 0); -x_114 = lean_ctor_get(x_106, 1); -lean_dec(x_114); -lean_inc(x_103); -x_115 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); -if (x_115 == 0) -{ -lean_object* x_116; +x_107 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_Dmain(x_1, x_2, x_103, x_4, x_5); +x_108 = lean_ctor_get(x_107, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_108, 1); +lean_inc(x_109); +x_110 = lean_unbox(x_109); +if (x_110 == 0) +{ +uint8_t x_111; +x_111 = !lean_is_exclusive(x_107); +if (x_111 == 0) +{ +lean_object* x_112; lean_object* x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_107, 1); +x_113 = lean_ctor_get(x_107, 0); +lean_dec(x_113); +x_114 = !lean_is_exclusive(x_108); +if (x_114 == 0) +{ +lean_object* x_115; lean_object* x_116; uint8_t x_117; +x_115 = lean_ctor_get(x_108, 0); +x_116 = lean_ctor_get(x_108, 1); +lean_dec(x_116); +lean_inc(x_105); +x_117 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_105); +if (x_117 == 0) +{ +lean_object* x_118; lean_dec(x_4); lean_dec(x_1); -x_116 = l_Lean_IR_FnBody_setBody(x_103, x_113); -lean_ctor_set(x_106, 0, x_116); -return x_105; +x_118 = l_Lean_IR_FnBody_setBody(x_105, x_115); +lean_ctor_set(x_108, 0, x_118); +return x_107; } else { -lean_object* x_117; uint8_t x_118; -lean_free_object(x_105); -lean_dec(x_107); -x_117 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_113, x_4, x_110); +lean_object* x_119; uint8_t x_120; +lean_free_object(x_107); +lean_dec(x_109); +x_119 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_115, x_4, x_112); lean_dec(x_4); -x_118 = !lean_is_exclusive(x_117); -if (x_118 == 0) -{ -lean_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; -x_119 = lean_ctor_get(x_117, 0); -x_120 = lean_ctor_get(x_117, 1); -x_121 = l_Lean_IR_FnBody_setBody(x_103, x_119); -x_122 = 1; -x_123 = lean_box(x_122); -lean_ctor_set(x_117, 1, x_123); -lean_ctor_set(x_117, 0, x_121); -lean_ctor_set(x_106, 1, x_120); -lean_ctor_set(x_106, 0, x_117); -return x_106; -} -else -{ -lean_object* x_124; lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; -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 = l_Lean_IR_FnBody_setBody(x_103, x_124); -x_127 = 1; -x_128 = lean_box(x_127); -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_126); -lean_ctor_set(x_129, 1, x_128); -lean_ctor_set(x_106, 1, x_125); -lean_ctor_set(x_106, 0, x_129); -return x_106; -} -} -} -else -{ -lean_object* x_130; uint8_t x_131; -x_130 = lean_ctor_get(x_106, 0); -lean_inc(x_130); -lean_dec(x_106); -lean_inc(x_103); -x_131 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); -if (x_131 == 0) -{ -lean_object* x_132; lean_object* x_133; +x_120 = !lean_is_exclusive(x_119); +if (x_120 == 0) +{ +lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; lean_object* x_125; +x_121 = lean_ctor_get(x_119, 0); +x_122 = lean_ctor_get(x_119, 1); +x_123 = l_Lean_IR_FnBody_setBody(x_105, x_121); +x_124 = 1; +x_125 = lean_box(x_124); +lean_ctor_set(x_119, 1, x_125); +lean_ctor_set(x_119, 0, x_123); +lean_ctor_set(x_108, 1, x_122); +lean_ctor_set(x_108, 0, x_119); +return x_108; +} +else +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; +x_126 = lean_ctor_get(x_119, 0); +x_127 = lean_ctor_get(x_119, 1); +lean_inc(x_127); +lean_inc(x_126); +lean_dec(x_119); +x_128 = l_Lean_IR_FnBody_setBody(x_105, x_126); +x_129 = 1; +x_130 = lean_box(x_129); +x_131 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_131, 0, x_128); +lean_ctor_set(x_131, 1, x_130); +lean_ctor_set(x_108, 1, x_127); +lean_ctor_set(x_108, 0, x_131); +return x_108; +} +} +} +else +{ +lean_object* x_132; uint8_t x_133; +x_132 = lean_ctor_get(x_108, 0); +lean_inc(x_132); +lean_dec(x_108); +lean_inc(x_105); +x_133 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_105); +if (x_133 == 0) +{ +lean_object* x_134; lean_object* x_135; lean_dec(x_4); lean_dec(x_1); -x_132 = l_Lean_IR_FnBody_setBody(x_103, x_130); -x_133 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_107); -lean_ctor_set(x_105, 0, x_133); -return x_105; +x_134 = l_Lean_IR_FnBody_setBody(x_105, x_132); +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_134); +lean_ctor_set(x_135, 1, x_109); +lean_ctor_set(x_107, 0, x_135); +return x_107; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; uint8_t x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; -lean_free_object(x_105); -lean_dec(x_107); -x_134 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_130, x_4, x_110); +lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +lean_free_object(x_107); +lean_dec(x_109); +x_136 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_132, x_4, x_112); lean_dec(x_4); -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_137 = x_134; +x_137 = lean_ctor_get(x_136, 0); +lean_inc(x_137); +x_138 = lean_ctor_get(x_136, 1); +lean_inc(x_138); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_139 = x_136; } else { - lean_dec_ref(x_134); - x_137 = lean_box(0); -} -x_138 = l_Lean_IR_FnBody_setBody(x_103, x_135); -x_139 = 1; -x_140 = lean_box(x_139); -if (lean_is_scalar(x_137)) { - x_141 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_136); + x_139 = lean_box(0); +} +x_140 = l_Lean_IR_FnBody_setBody(x_105, x_137); +x_141 = 1; +x_142 = lean_box(x_141); +if (lean_is_scalar(x_139)) { + x_143 = lean_alloc_ctor(0, 2, 0); } else { - x_141 = x_137; + x_143 = x_139; } -lean_ctor_set(x_141, 0, x_138); -lean_ctor_set(x_141, 1, x_140); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_141); -lean_ctor_set(x_142, 1, x_136); -return x_142; +lean_ctor_set(x_143, 0, x_140); +lean_ctor_set(x_143, 1, x_142); +x_144 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_144, 0, x_143); +lean_ctor_set(x_144, 1, x_138); +return x_144; } } } else { -lean_object* x_143; lean_object* x_144; lean_object* x_145; uint8_t x_146; -x_143 = lean_ctor_get(x_105, 1); -lean_inc(x_143); -lean_dec(x_105); -x_144 = lean_ctor_get(x_106, 0); -lean_inc(x_144); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_145 = x_106; +lean_object* x_145; lean_object* x_146; lean_object* x_147; uint8_t x_148; +x_145 = lean_ctor_get(x_107, 1); +lean_inc(x_145); +lean_dec(x_107); +x_146 = lean_ctor_get(x_108, 0); +lean_inc(x_146); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_147 = x_108; } else { - lean_dec_ref(x_106); - x_145 = lean_box(0); + lean_dec_ref(x_108); + x_147 = lean_box(0); } -lean_inc(x_103); -x_146 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_103); -if (x_146 == 0) +lean_inc(x_105); +x_148 = l_Lean_IR_HasIndex_visitFnBody(x_1, x_105); +if (x_148 == 0) { -lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_dec(x_4); lean_dec(x_1); -x_147 = l_Lean_IR_FnBody_setBody(x_103, x_144); -if (lean_is_scalar(x_145)) { - x_148 = lean_alloc_ctor(0, 2, 0); +x_149 = l_Lean_IR_FnBody_setBody(x_105, x_146); +if (lean_is_scalar(x_147)) { + x_150 = lean_alloc_ctor(0, 2, 0); } else { - x_148 = x_145; + x_150 = x_147; } -lean_ctor_set(x_148, 0, x_147); -lean_ctor_set(x_148, 1, x_107); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_143); -return x_149; +lean_ctor_set(x_150, 0, x_149); +lean_ctor_set(x_150, 1, x_109); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_150); +lean_ctor_set(x_151, 1, x_145); +return x_151; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; uint8_t x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -lean_dec(x_107); -x_150 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_144, x_4, x_143); +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; uint8_t x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +lean_dec(x_109); +x_152 = l___private_Lean_Compiler_IR_ResetReuse_0__Lean_IR_ResetReuse_tryS(x_1, x_2, x_146, x_4, x_145); lean_dec(x_4); -x_151 = lean_ctor_get(x_150, 0); -lean_inc(x_151); -x_152 = lean_ctor_get(x_150, 1); -lean_inc(x_152); -if (lean_is_exclusive(x_150)) { - lean_ctor_release(x_150, 0); - lean_ctor_release(x_150, 1); - x_153 = x_150; +x_153 = lean_ctor_get(x_152, 0); +lean_inc(x_153); +x_154 = lean_ctor_get(x_152, 1); +lean_inc(x_154); +if (lean_is_exclusive(x_152)) { + lean_ctor_release(x_152, 0); + lean_ctor_release(x_152, 1); + x_155 = x_152; } else { - lean_dec_ref(x_150); - x_153 = lean_box(0); -} -x_154 = l_Lean_IR_FnBody_setBody(x_103, x_151); -x_155 = 1; -x_156 = lean_box(x_155); -if (lean_is_scalar(x_153)) { - x_157 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_152); + x_155 = lean_box(0); +} +x_156 = l_Lean_IR_FnBody_setBody(x_105, x_153); +x_157 = 1; +x_158 = lean_box(x_157); +if (lean_is_scalar(x_155)) { + x_159 = lean_alloc_ctor(0, 2, 0); } else { - x_157 = x_153; + x_159 = x_155; } -lean_ctor_set(x_157, 0, x_154); -lean_ctor_set(x_157, 1, x_156); -if (lean_is_scalar(x_145)) { - x_158 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_159, 0, x_156); +lean_ctor_set(x_159, 1, x_158); +if (lean_is_scalar(x_147)) { + x_160 = lean_alloc_ctor(0, 2, 0); } else { - x_158 = x_145; + x_160 = x_147; } -lean_ctor_set(x_158, 0, x_157); -lean_ctor_set(x_158, 1, x_152); -return x_158; +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_154); +return x_160; } } } else { -uint8_t x_159; +uint8_t x_161; lean_dec(x_4); lean_dec(x_1); -x_159 = !lean_is_exclusive(x_105); -if (x_159 == 0) -{ -lean_object* x_160; uint8_t x_161; -x_160 = lean_ctor_get(x_105, 0); -lean_dec(x_160); -x_161 = !lean_is_exclusive(x_106); +x_161 = !lean_is_exclusive(x_107); if (x_161 == 0) { -lean_object* x_162; lean_object* x_163; lean_object* x_164; -x_162 = lean_ctor_get(x_106, 0); -x_163 = lean_ctor_get(x_106, 1); -lean_dec(x_163); -x_164 = l_Lean_IR_FnBody_setBody(x_103, x_162); -lean_ctor_set(x_106, 0, x_164); -return x_105; +lean_object* x_162; uint8_t x_163; +x_162 = lean_ctor_get(x_107, 0); +lean_dec(x_162); +x_163 = !lean_is_exclusive(x_108); +if (x_163 == 0) +{ +lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_164 = lean_ctor_get(x_108, 0); +x_165 = lean_ctor_get(x_108, 1); +lean_dec(x_165); +x_166 = l_Lean_IR_FnBody_setBody(x_105, x_164); +lean_ctor_set(x_108, 0, x_166); +return x_107; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_106, 0); -lean_inc(x_165); -lean_dec(x_106); -x_166 = l_Lean_IR_FnBody_setBody(x_103, x_165); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_166); -lean_ctor_set(x_167, 1, x_107); -lean_ctor_set(x_105, 0, x_167); -return x_105; +lean_object* x_167; lean_object* x_168; lean_object* x_169; +x_167 = lean_ctor_get(x_108, 0); +lean_inc(x_167); +lean_dec(x_108); +x_168 = l_Lean_IR_FnBody_setBody(x_105, x_167); +x_169 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_169, 0, x_168); +lean_ctor_set(x_169, 1, x_109); +lean_ctor_set(x_107, 0, x_169); +return x_107; } } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_168 = lean_ctor_get(x_105, 1); -lean_inc(x_168); -lean_dec(x_105); -x_169 = lean_ctor_get(x_106, 0); -lean_inc(x_169); -if (lean_is_exclusive(x_106)) { - lean_ctor_release(x_106, 0); - lean_ctor_release(x_106, 1); - x_170 = x_106; +lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_170 = lean_ctor_get(x_107, 1); +lean_inc(x_170); +lean_dec(x_107); +x_171 = lean_ctor_get(x_108, 0); +lean_inc(x_171); +if (lean_is_exclusive(x_108)) { + lean_ctor_release(x_108, 0); + lean_ctor_release(x_108, 1); + x_172 = x_108; } else { - lean_dec_ref(x_106); - x_170 = lean_box(0); + lean_dec_ref(x_108); + x_172 = lean_box(0); } -x_171 = l_Lean_IR_FnBody_setBody(x_103, x_169); -if (lean_is_scalar(x_170)) { - x_172 = lean_alloc_ctor(0, 2, 0); +x_173 = l_Lean_IR_FnBody_setBody(x_105, x_171); +if (lean_is_scalar(x_172)) { + x_174 = lean_alloc_ctor(0, 2, 0); } else { - x_172 = x_170; + x_174 = x_172; } -lean_ctor_set(x_172, 0, x_171); -lean_ctor_set(x_172, 1, x_107); -x_173 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_173, 0, x_172); -lean_ctor_set(x_173, 1, x_168); -return x_173; +lean_ctor_set(x_174, 0, x_173); +lean_ctor_set(x_174, 1, x_109); +x_175 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_175, 0, x_174); +lean_ctor_set(x_175, 1, x_170); +return x_175; } } } else { -uint8_t x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; +uint8_t x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; +lean_dec(x_105); lean_dec(x_103); -lean_dec(x_101); lean_dec(x_4); lean_dec(x_1); -x_174 = 1; -x_175 = lean_box(x_174); -x_176 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_176, 0, x_3); -lean_ctor_set(x_176, 1, x_175); -x_177 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_177, 0, x_176); -lean_ctor_set(x_177, 1, x_5); -return x_177; +x_176 = 1; +x_177 = lean_box(x_176); +x_178 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_178, 0, x_3); +lean_ctor_set(x_178, 1, x_177); +x_179 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_179, 0, x_178); +lean_ctor_set(x_179, 1, x_5); +return x_179; } } else { -lean_object* x_178; lean_object* x_179; uint8_t x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -x_178 = lean_ctor_get(x_4, 0); -lean_inc(x_178); +lean_object* x_180; lean_object* x_181; uint8_t x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +x_180 = lean_ctor_get(x_4, 0); +lean_inc(x_180); lean_dec(x_4); lean_inc(x_3); -x_179 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_178, x_1); +x_181 = l_Lean_IR_FnBody_hasLiveVar(x_3, x_180, x_1); lean_dec(x_1); -x_180 = lean_unbox(x_179); -lean_dec(x_179); -x_181 = lean_box(x_180); -x_182 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_182, 0, x_3); -lean_ctor_set(x_182, 1, x_181); -x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_5); -return x_183; +x_182 = lean_unbox(x_181); +lean_dec(x_181); +x_183 = lean_box(x_182); +x_184 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_184, 0, x_3); +lean_ctor_set(x_184, 1, x_183); +x_185 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_185, 0, x_184); +lean_ctor_set(x_185, 1, x_5); +return x_185; } } } @@ -2792,326 +2849,334 @@ return x_20; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; x_21 = lean_ctor_get(x_2, 0); x_22 = lean_ctor_get(x_2, 1); +x_23 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); lean_inc(x_22); lean_inc(x_21); lean_dec(x_2); lean_inc(x_10); lean_inc(x_6); lean_inc(x_5); -x_23 = l_Lean_IR_LocalContext_addJP(x_21, x_5, x_6, x_10); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -x_25 = l_Lean_IR_ResetReuse_R(x_8, x_24, x_11); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +x_24 = l_Lean_IR_LocalContext_addJP(x_21, x_5, x_6, x_10); +x_25 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_22); +lean_ctor_set_uint8(x_25, sizeof(void*)*2, x_23); +x_26 = l_Lean_IR_ResetReuse_R(x_8, x_25, x_11); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -if (lean_is_exclusive(x_25)) { - lean_ctor_release(x_25, 0); - lean_ctor_release(x_25, 1); - x_28 = x_25; +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_29 = x_26; } else { - lean_dec_ref(x_25); - x_28 = lean_box(0); + lean_dec_ref(x_26); + x_29 = lean_box(0); } -lean_ctor_set(x_1, 3, x_26); +lean_ctor_set(x_1, 3, x_27); lean_ctor_set(x_1, 2, x_10); -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_29)) { + x_30 = lean_alloc_ctor(0, 2, 0); } else { - x_29 = x_28; + x_30 = x_29; } -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_27); -return x_29; +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_28); +return x_30; } } 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; lean_object* x_47; -x_30 = lean_ctor_get(x_1, 0); -x_31 = lean_ctor_get(x_1, 1); -x_32 = lean_ctor_get(x_1, 2); -x_33 = lean_ctor_get(x_1, 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; 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_31 = lean_ctor_get(x_1, 0); +x_32 = lean_ctor_get(x_1, 1); +x_33 = lean_ctor_get(x_1, 2); +x_34 = lean_ctor_get(x_1, 3); +lean_inc(x_34); lean_inc(x_33); lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); lean_dec(x_1); lean_inc(x_2); -x_34 = l_Lean_IR_ResetReuse_R(x_32, x_2, x_3); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); +x_35 = l_Lean_IR_ResetReuse_R(x_33, x_2, x_3); +x_36 = lean_ctor_get(x_35, 0); lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_ctor_get(x_2, 0); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); -x_38 = lean_ctor_get(x_2, 1); +lean_dec(x_35); +x_38 = lean_ctor_get(x_2, 0); lean_inc(x_38); +x_39 = lean_ctor_get(x_2, 1); +lean_inc(x_39); +x_40 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); if (lean_is_exclusive(x_2)) { lean_ctor_release(x_2, 0); lean_ctor_release(x_2, 1); - x_39 = x_2; + x_41 = x_2; } else { lean_dec_ref(x_2); - x_39 = lean_box(0); + x_41 = lean_box(0); } -lean_inc(x_35); +lean_inc(x_36); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -x_40 = l_Lean_IR_LocalContext_addJP(x_37, x_30, x_31, x_35); -if (lean_is_scalar(x_39)) { - x_41 = lean_alloc_ctor(0, 2, 0); +x_42 = l_Lean_IR_LocalContext_addJP(x_38, x_31, x_32, x_36); +if (lean_is_scalar(x_41)) { + x_43 = lean_alloc_ctor(0, 2, 1); } else { - x_41 = x_39; -} -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_38); -x_42 = l_Lean_IR_ResetReuse_R(x_33, x_41, x_36); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_45 = x_42; + x_43 = x_41; +} +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_39); +lean_ctor_set_uint8(x_43, sizeof(void*)*2, x_40); +x_44 = l_Lean_IR_ResetReuse_R(x_34, x_43, 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); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_47 = x_44; } else { - lean_dec_ref(x_42); - x_45 = lean_box(0); -} -x_46 = lean_alloc_ctor(1, 4, 0); -lean_ctor_set(x_46, 0, x_30); -lean_ctor_set(x_46, 1, x_31); -lean_ctor_set(x_46, 2, x_35); -lean_ctor_set(x_46, 3, x_43); -if (lean_is_scalar(x_45)) { - x_47 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_44); + x_47 = lean_box(0); +} +x_48 = lean_alloc_ctor(1, 4, 0); +lean_ctor_set(x_48, 0, x_31); +lean_ctor_set(x_48, 1, x_32); +lean_ctor_set(x_48, 2, x_36); +lean_ctor_set(x_48, 3, x_45); +if (lean_is_scalar(x_47)) { + x_49 = lean_alloc_ctor(0, 2, 0); } else { - x_47 = x_45; + x_49 = x_47; } -lean_ctor_set(x_47, 0, x_46); -lean_ctor_set(x_47, 1, x_44); -return x_47; +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_46); +return x_49; } } case 10: { -uint8_t x_48; -x_48 = !lean_is_exclusive(x_1); -if (x_48 == 0) +uint8_t x_50; +x_50 = !lean_is_exclusive(x_1); +if (x_50 == 0) { -uint8_t x_49; -x_49 = !lean_is_exclusive(x_2); -if (x_49 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_2); +if (x_51 == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; size_t x_55; size_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; -x_50 = lean_ctor_get(x_1, 1); -x_51 = lean_ctor_get(x_1, 3); -x_52 = lean_ctor_get(x_2, 1); -lean_inc(x_50); +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; lean_object* x_56; size_t x_57; size_t x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; +x_52 = lean_ctor_get(x_1, 1); +x_53 = lean_ctor_get(x_1, 3); +x_54 = lean_ctor_get(x_2, 1); lean_inc(x_52); -x_53 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_52, x_50); -x_54 = lean_array_get_size(x_51); -x_55 = lean_usize_of_nat(x_54); -lean_dec(x_54); -x_56 = 0; -x_57 = lean_box(0); -lean_inc(x_50); -x_58 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_52, x_50, x_57); -lean_ctor_set(x_2, 1, x_58); -lean_inc(x_50); -x_59 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_50, x_53, x_55, x_56, x_51, x_2, x_3); -x_60 = !lean_is_exclusive(x_59); -if (x_60 == 0) +lean_inc(x_54); +x_55 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_54, x_52); +x_56 = lean_array_get_size(x_53); +x_57 = lean_usize_of_nat(x_56); +lean_dec(x_56); +x_58 = 0; +x_59 = lean_box(0); +lean_inc(x_52); +x_60 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_54, x_52, x_59); +lean_ctor_set(x_2, 1, x_60); +lean_inc(x_52); +x_61 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_52, x_55, x_57, x_58, x_53, x_2, x_3); +x_62 = !lean_is_exclusive(x_61); +if (x_62 == 0) { -lean_object* x_61; -x_61 = lean_ctor_get(x_59, 0); -lean_ctor_set(x_1, 3, x_61); -lean_ctor_set(x_59, 0, x_1); -return x_59; +lean_object* x_63; +x_63 = lean_ctor_get(x_61, 0); +lean_ctor_set(x_1, 3, x_63); +lean_ctor_set(x_61, 0, x_1); +return x_61; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_59, 0); -x_63 = lean_ctor_get(x_59, 1); -lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_59); -lean_ctor_set(x_1, 3, x_62); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_1); -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_61, 0); +x_65 = lean_ctor_get(x_61, 1); +lean_inc(x_65); +lean_inc(x_64); +lean_dec(x_61); +lean_ctor_set(x_1, 3, x_64); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_1); +lean_ctor_set(x_66, 1, x_65); +return x_66; } } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; lean_object* x_70; size_t x_71; size_t 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_65 = lean_ctor_get(x_1, 1); -x_66 = lean_ctor_get(x_1, 3); -x_67 = lean_ctor_get(x_2, 0); -x_68 = lean_ctor_get(x_2, 1); -lean_inc(x_68); -lean_inc(x_67); +lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; uint8_t x_72; lean_object* x_73; size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_67 = lean_ctor_get(x_1, 1); +x_68 = lean_ctor_get(x_1, 3); +x_69 = lean_ctor_get(x_2, 0); +x_70 = lean_ctor_get(x_2, 1); +x_71 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); +lean_inc(x_70); +lean_inc(x_69); lean_dec(x_2); -lean_inc(x_65); -lean_inc(x_68); -x_69 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_68, x_65); -x_70 = lean_array_get_size(x_66); -x_71 = lean_usize_of_nat(x_70); -lean_dec(x_70); -x_72 = 0; -x_73 = lean_box(0); -lean_inc(x_65); -x_74 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_68, x_65, x_73); -x_75 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_75, 0, x_67); -lean_ctor_set(x_75, 1, x_74); -lean_inc(x_65); -x_76 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_65, x_69, x_71, x_72, x_66, x_75, x_3); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_76, 1); -lean_inc(x_78); -if (lean_is_exclusive(x_76)) { - lean_ctor_release(x_76, 0); - lean_ctor_release(x_76, 1); - x_79 = x_76; +lean_inc(x_67); +lean_inc(x_70); +x_72 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_70, x_67); +x_73 = lean_array_get_size(x_68); +x_74 = lean_usize_of_nat(x_73); +lean_dec(x_73); +x_75 = 0; +x_76 = lean_box(0); +lean_inc(x_67); +x_77 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_70, x_67, x_76); +x_78 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_78, 0, x_69); +lean_ctor_set(x_78, 1, x_77); +lean_ctor_set_uint8(x_78, sizeof(void*)*2, x_71); +lean_inc(x_67); +x_79 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_67, x_72, x_74, x_75, x_68, x_78, x_3); +x_80 = lean_ctor_get(x_79, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_79, 1); +lean_inc(x_81); +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_76); - x_79 = lean_box(0); + lean_dec_ref(x_79); + x_82 = lean_box(0); } -lean_ctor_set(x_1, 3, x_77); -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_1, 3, x_80); +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(0, 2, 0); } else { - x_80 = x_79; + x_83 = x_82; } -lean_ctor_set(x_80, 0, x_1); -lean_ctor_set(x_80, 1, x_78); -return x_80; +lean_ctor_set(x_83, 0, x_1); +lean_ctor_set(x_83, 1, x_81); +return x_83; } } 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; uint8_t x_88; lean_object* x_89; size_t x_90; size_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_81 = lean_ctor_get(x_1, 0); -x_82 = lean_ctor_get(x_1, 1); -x_83 = lean_ctor_get(x_1, 2); -x_84 = lean_ctor_get(x_1, 3); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; lean_object* x_91; uint8_t x_92; lean_object* x_93; size_t x_94; size_t x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_84 = lean_ctor_get(x_1, 0); +x_85 = lean_ctor_get(x_1, 1); +x_86 = lean_ctor_get(x_1, 2); +x_87 = lean_ctor_get(x_1, 3); +lean_inc(x_87); +lean_inc(x_86); +lean_inc(x_85); lean_inc(x_84); -lean_inc(x_83); -lean_inc(x_82); -lean_inc(x_81); lean_dec(x_1); -x_85 = lean_ctor_get(x_2, 0); -lean_inc(x_85); -x_86 = lean_ctor_get(x_2, 1); -lean_inc(x_86); +x_88 = lean_ctor_get(x_2, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_2, 1); +lean_inc(x_89); +x_90 = lean_ctor_get_uint8(x_2, sizeof(void*)*2); if (lean_is_exclusive(x_2)) { lean_ctor_release(x_2, 0); lean_ctor_release(x_2, 1); - x_87 = x_2; + x_91 = x_2; } else { lean_dec_ref(x_2); - x_87 = lean_box(0); + x_91 = lean_box(0); } -lean_inc(x_82); -lean_inc(x_86); -x_88 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_86, x_82); -x_89 = lean_array_get_size(x_84); -x_90 = lean_usize_of_nat(x_89); -lean_dec(x_89); -x_91 = 0; -x_92 = lean_box(0); -lean_inc(x_82); -x_93 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_86, x_82, x_92); -if (lean_is_scalar(x_87)) { - x_94 = lean_alloc_ctor(0, 2, 0); +lean_inc(x_85); +lean_inc(x_89); +x_92 = l_Lean_PersistentHashMap_contains___at_Lean_IR_ResetReuse_R___spec__1(x_89, x_85); +x_93 = lean_array_get_size(x_87); +x_94 = lean_usize_of_nat(x_93); +lean_dec(x_93); +x_95 = 0; +x_96 = lean_box(0); +lean_inc(x_85); +x_97 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_89, x_85, x_96); +if (lean_is_scalar(x_91)) { + x_98 = lean_alloc_ctor(0, 2, 1); } else { - x_94 = x_87; -} -lean_ctor_set(x_94, 0, x_85); -lean_ctor_set(x_94, 1, x_93); -lean_inc(x_82); -x_95 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_82, x_88, x_90, x_91, x_84, x_94, x_3); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; + x_98 = x_91; +} +lean_ctor_set(x_98, 0, x_88); +lean_ctor_set(x_98, 1, x_97); +lean_ctor_set_uint8(x_98, sizeof(void*)*2, x_90); +lean_inc(x_85); +x_99 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_85, x_92, x_94, x_95, x_87, x_98, x_3); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_102 = x_99; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); -} -x_99 = lean_alloc_ctor(10, 4, 0); -lean_ctor_set(x_99, 0, x_81); -lean_ctor_set(x_99, 1, x_82); -lean_ctor_set(x_99, 2, x_83); -lean_ctor_set(x_99, 3, x_96); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_99); + x_102 = lean_box(0); +} +x_103 = lean_alloc_ctor(10, 4, 0); +lean_ctor_set(x_103, 0, x_84); +lean_ctor_set(x_103, 1, x_85); +lean_ctor_set(x_103, 2, x_86); +lean_ctor_set(x_103, 3, x_100); +if (lean_is_scalar(x_102)) { + x_104 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_104 = x_102; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_104, 0, x_103); +lean_ctor_set(x_104, 1, x_101); +return x_104; } } default: { -uint8_t x_101; -x_101 = l_Lean_IR_FnBody_isTerminal(x_1); -if (x_101 == 0) +uint8_t x_105; +x_105 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_105 == 0) { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; uint8_t x_106; -x_102 = l_Lean_IR_FnBody_body(x_1); -x_103 = lean_box(13); -x_104 = l_Lean_IR_FnBody_setBody(x_1, x_103); -x_105 = l_Lean_IR_ResetReuse_R(x_102, x_2, x_3); -x_106 = !lean_is_exclusive(x_105); -if (x_106 == 0) +lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; uint8_t x_110; +x_106 = l_Lean_IR_FnBody_body(x_1); +x_107 = lean_box(13); +x_108 = l_Lean_IR_FnBody_setBody(x_1, x_107); +x_109 = l_Lean_IR_ResetReuse_R(x_106, x_2, x_3); +x_110 = !lean_is_exclusive(x_109); +if (x_110 == 0) { -lean_object* x_107; lean_object* x_108; -x_107 = lean_ctor_get(x_105, 0); -x_108 = l_Lean_IR_FnBody_setBody(x_104, x_107); -lean_ctor_set(x_105, 0, x_108); -return x_105; +lean_object* x_111; lean_object* x_112; +x_111 = lean_ctor_get(x_109, 0); +x_112 = l_Lean_IR_FnBody_setBody(x_108, x_111); +lean_ctor_set(x_109, 0, x_112); +return x_109; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = lean_ctor_get(x_105, 0); -x_110 = lean_ctor_get(x_105, 1); -lean_inc(x_110); -lean_inc(x_109); -lean_dec(x_105); -x_111 = l_Lean_IR_FnBody_setBody(x_104, x_109); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_111); -lean_ctor_set(x_112, 1, x_110); -return x_112; +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_113 = lean_ctor_get(x_109, 0); +x_114 = lean_ctor_get(x_109, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_109); +x_115 = l_Lean_IR_FnBody_setBody(x_108, x_113); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +return x_116; } } else { -lean_object* x_113; +lean_object* x_117; lean_dec(x_2); -x_113 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_113, 0, x_1); -lean_ctor_set(x_113, 1, x_3); -return x_113; +x_117 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_117, 0, x_1); +lean_ctor_set(x_117, 1, x_3); +return x_117; } } } @@ -3188,39 +3253,243 @@ x_11 = l_Array_mapMUnsafe_map___at_Lean_IR_ResetReuse_R___spec__8(x_1, x_8, x_9, return x_11; } } -static lean_object* _init_l_Lean_IR_Decl_insertResetReuse___closed__1() { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_box(0); -x_2 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1; -x_3 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_3, 0, x_1); -lean_ctor_set(x_3, 1, x_2); -return x_3; +uint8_t x_6; +x_6 = lean_usize_dec_eq(x_2, 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; size_t x_12; size_t x_13; +lean_dec(x_4); +x_7 = lean_array_uget(x_1, x_2); +x_8 = l_Lean_IR_AltCore_body(x_7); +lean_dec(x_7); +x_9 = l_Lean_IR_ResetReuse_collectResets(x_8, x_5); +x_10 = lean_ctor_get(x_9, 0); +lean_inc(x_10); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); +lean_dec(x_9); +x_12 = 1; +x_13 = lean_usize_add(x_2, x_12); +x_2 = x_13; +x_4 = x_10; +x_5 = x_11; +goto _start; +} +else +{ +lean_object* x_15; +x_15 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_15, 0, x_4); +lean_ctor_set(x_15, 1, x_5); +return x_15; } } -LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object* x_1) { +} +LEAN_EXPORT lean_object* l_Lean_IR_ResetReuse_collectResets(lean_object* x_1, lean_object* x_2) { +_start: +{ +switch (lean_obj_tag(x_1)) { +case 0: +{ +lean_object* x_3; +x_3 = lean_ctor_get(x_1, 2); +lean_inc(x_3); +if (lean_obj_tag(x_3) == 1) +{ +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_4 = lean_ctor_get(x_1, 3); +lean_inc(x_4); +lean_dec(x_1); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_box(0); +x_7 = l_Lean_PersistentHashMap_insert___at_Lean_IR_ResetReuse_R___spec__4(x_2, x_5, x_6); +x_1 = x_4; +x_2 = x_7; +goto _start; +} +else +{ +uint8_t x_9; +lean_dec(x_3); +x_9 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_9 == 0) +{ +lean_object* x_10; +x_10 = l_Lean_IR_FnBody_body(x_1); +lean_dec(x_1); +x_1 = x_10; +goto _start; +} +else +{ +lean_object* x_12; lean_object* x_13; +lean_dec(x_1); +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_2); +return x_13; +} +} +} +case 1: +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_ctor_get(x_1, 2); +lean_inc(x_14); +x_15 = lean_ctor_get(x_1, 3); +lean_inc(x_15); +lean_dec(x_1); +x_16 = l_Lean_IR_ResetReuse_collectResets(x_14, x_2); +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_1 = x_15; +x_2 = x_17; +goto _start; +} +case 10: +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_19 = lean_ctor_get(x_1, 3); +lean_inc(x_19); +lean_dec(x_1); +x_20 = lean_array_get_size(x_19); +x_21 = lean_unsigned_to_nat(0u); +x_22 = lean_nat_dec_lt(x_21, x_20); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; +lean_dec(x_20); +lean_dec(x_19); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_24, 1, x_2); +return x_24; +} +else +{ +uint8_t x_25; +x_25 = lean_nat_dec_le(x_20, x_20); +if (x_25 == 0) +{ +lean_object* x_26; lean_object* x_27; +lean_dec(x_20); +lean_dec(x_19); +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_2); +return x_27; +} +else +{ +size_t x_28; size_t x_29; lean_object* x_30; lean_object* x_31; +x_28 = 0; +x_29 = lean_usize_of_nat(x_20); +lean_dec(x_20); +x_30 = lean_box(0); +x_31 = l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___spec__1(x_19, x_28, x_29, x_30, x_2); +lean_dec(x_19); +return x_31; +} +} +} +default: +{ +uint8_t x_32; +x_32 = l_Lean_IR_FnBody_isTerminal(x_1); +if (x_32 == 0) +{ +lean_object* x_33; +x_33 = l_Lean_IR_FnBody_body(x_1); +lean_dec(x_1); +x_1 = x_33; +goto _start; +} +else +{ +lean_object* x_35; lean_object* x_36; +lean_dec(x_1); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_2); +return x_36; +} +} +} +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___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_2); +lean_dec(x_2); +x_7 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_8 = l_Array_foldlMUnsafe_fold___at_Lean_IR_ResetReuse_collectResets___spec__1(x_1, x_6, x_7, x_4, x_5); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuseCore(lean_object* x_1, uint8_t x_2) { _start: { if (lean_obj_tag(x_1) == 0) { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_2 = lean_ctor_get(x_1, 3); -lean_inc(x_2); -x_3 = lean_unsigned_to_nat(0u); +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_3 = lean_ctor_get(x_1, 3); +lean_inc(x_3); +x_4 = lean_unsigned_to_nat(0u); lean_inc(x_1); -x_4 = l_Lean_IR_MaxIndex_collectDecl(x_1, x_3); -x_5 = lean_unsigned_to_nat(1u); -x_6 = lean_nat_add(x_4, x_5); -lean_dec(x_4); -x_7 = l_Lean_IR_Decl_insertResetReuse___closed__1; -x_8 = l_Lean_IR_ResetReuse_R(x_2, x_7, x_6); -x_9 = lean_ctor_get(x_8, 0); -lean_inc(x_9); -lean_dec(x_8); -x_10 = l_Lean_IR_Decl_updateBody_x21(x_1, x_9); -return x_10; +x_5 = l_Lean_IR_MaxIndex_collectDecl(x_1, x_4); +x_6 = lean_unsigned_to_nat(1u); +x_7 = lean_nat_add(x_5, x_6); +lean_dec(x_5); +x_8 = lean_box(0); +if (x_2 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_9 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1; +x_10 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_10, 0, x_8); +lean_ctor_set(x_10, 1, x_9); +lean_ctor_set_uint8(x_10, sizeof(void*)*2, x_2); +x_11 = l_Lean_IR_ResetReuse_R(x_3, x_10, x_7); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_IR_Decl_updateBody_x21(x_1, x_12); +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; +x_14 = l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1; +lean_inc(x_3); +x_15 = l_Lean_IR_ResetReuse_collectResets(x_3, x_14); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_17, 0, x_8); +lean_ctor_set(x_17, 1, x_16); +lean_ctor_set_uint8(x_17, sizeof(void*)*2, x_2); +x_18 = l_Lean_IR_ResetReuse_R(x_3, x_17, x_7); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +lean_dec(x_18); +x_20 = l_Lean_IR_Decl_updateBody_x21(x_1, x_19); +return x_20; +} } else { @@ -3228,6 +3497,27 @@ return x_1; } } } +LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuseCore___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_2); +lean_dec(x_2); +x_4 = l_Lean_IR_Decl_insertResetReuseCore(x_1, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_IR_Decl_insertResetReuse(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; +x_2 = 0; +x_3 = l_Lean_IR_Decl_insertResetReuseCore(x_1, x_2); +x_4 = 1; +x_5 = l_Lean_IR_Decl_insertResetReuseCore(x_3, x_4); +return x_5; +} +} lean_object* initialize_Lean_Compiler_IR_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_IR_LiveVars(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Compiler_IR_Format(uint8_t builtin, lean_object*); @@ -3247,26 +3537,25 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_IR_ResetReuse_Context_lctx___default = _init_l_Lean_IR_ResetReuse_Context_lctx___default(); lean_mark_persistent(l_Lean_IR_ResetReuse_Context_lctx___default); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__1); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__2); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__3); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__4); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1___closed__5); -l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1(); -lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_casesVars___default___spec__1); -l_Lean_IR_ResetReuse_Context_casesVars___default = _init_l_Lean_IR_ResetReuse_Context_casesVars___default(); -lean_mark_persistent(l_Lean_IR_ResetReuse_Context_casesVars___default); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__1); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__2); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__3 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__3(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__3); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__4 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__4(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__4); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1___closed__5); +l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1 = _init_l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1(); +lean_mark_persistent(l_Lean_PersistentHashMap_empty___at_Lean_IR_ResetReuse_Context_alreadyFound___default___spec__1); +l_Lean_IR_ResetReuse_Context_alreadyFound___default = _init_l_Lean_IR_ResetReuse_Context_alreadyFound___default(); +lean_mark_persistent(l_Lean_IR_ResetReuse_Context_alreadyFound___default); +l_Lean_IR_ResetReuse_Context_relaxedReuse___default = _init_l_Lean_IR_ResetReuse_Context_relaxedReuse___default(); l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1 = _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__1(); l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2 = _init_l_Lean_PersistentHashMap_containsAux___at_Lean_IR_ResetReuse_R___spec__2___closed__2(); l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1 = _init_l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1(); lean_mark_persistent(l_Lean_PersistentHashMap_insertAux___at_Lean_IR_ResetReuse_R___spec__5___closed__1); -l_Lean_IR_Decl_insertResetReuse___closed__1 = _init_l_Lean_IR_Decl_insertResetReuse___closed__1(); -lean_mark_persistent(l_Lean_IR_Decl_insertResetReuse___closed__1); return lean_io_result_mk_ok(lean_box(0)); } #ifdef __cplusplus diff --git a/stage0/stdlib/Lean/Compiler/InitAttr.c b/stage0/stdlib/Lean/Compiler/InitAttr.c index dd5a2a09bcc8..f732d506affa 100644 --- a/stage0/stdlib/Lean/Compiler/InitAttr.c +++ b/stage0/stdlib/Lean/Compiler/InitAttr.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Compiler.InitAttr -// Imports: Lean.Elab.InfoTree.Main +// Imports: Lean.AddDecl Lean.Elab.InfoTree.Main #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -14,12 +14,9 @@ extern "C" { #endif lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__13; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___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_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180_(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__29; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); @@ -37,7 +34,6 @@ LEAN_EXPORT lean_object* l_Lean_setBuiltinInitAttr(lean_object*, lean_object*, l static lean_object* l_Lean_declareBuiltin___closed__13; LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getInitFnNameForCore_x3f(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4; 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_interpretedModInits; @@ -49,11 +45,9 @@ static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed_ static lean_object* l_Lean_declareBuiltin___closed__5; LEAN_EXPORT lean_object* l_Lean_builtinInitAttr; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4; lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_registerInitAttrUnsafe___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isIOUnitInitFn___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__28; lean_object* lean_environment_find(lean_object*, lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___closed__1; @@ -88,10 +82,8 @@ uint8_t lean_expr_eqv(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_get_init_fn_name_for(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_is_io_unit_regular_init_fn(lean_object*, lean_object*); -lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_is_io_unit_builtin_init_fn(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208_(lean_object*); lean_object* lean_run_init(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerInitAttrUnsafe___closed__3; @@ -121,6 +113,7 @@ lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__12; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_registerInitAttrUnsafe___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__23; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_registerInitAttrUnsafe___spec__5___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_get_builtin_init_fn_name_for(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerInitAttr(lean_object*, uint8_t, lean_object*, lean_object*); @@ -131,8 +124,8 @@ LEAN_EXPORT uint8_t l___private_Lean_Compiler_InitAttr_0__Lean_isUnitType(lean_o lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1___closed__4; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__3; -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2; lean_object* l_Lean_ParametricAttribute_setParam___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150_(lean_object*); LEAN_EXPORT lean_object* l_Lean_runModInit___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__21; LEAN_EXPORT lean_object* l_Lean_regularInitAttr; @@ -151,9 +144,8 @@ LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__3___boxed(lean_ lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_ParametricAttribute_getParam_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2; LEAN_EXPORT lean_object* l_Lean_hasInitAttr___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3; uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); static lean_object* l_Lean_declareBuiltin___closed__11; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__18; @@ -163,11 +155,15 @@ static lean_object* l_Lean_declareBuiltin___closed__12; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__25; lean_object* l_Lean_registerParametricAttribute___rarg(lean_object*, lean_object*); static lean_object* l_Lean_declareBuiltin___closed__7; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178_(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__26; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__11; size_t lean_usize_add(size_t, size_t); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__1; extern lean_object* l_Lean_instInhabitedName; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___spec__7___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_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l___private_Lean_Compiler_InitAttr_0__Lean_isUnitType___boxed(lean_object*); @@ -175,11 +171,14 @@ lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_registerInitAttrUnsafe___spec__3(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_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3; LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Compiler_InitAttr_0__Lean_isIOUnit___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_registerInitAttrUnsafe___spec__6___closed__1; +lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__5; LEAN_EXPORT lean_object* l_Lean_registerInitAttrUnsafe___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_isIOUnitInitFn(lean_object*, lean_object*); @@ -194,6 +193,7 @@ lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_val static lean_object* l_Lean_declareBuiltin___closed__6; static lean_object* l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__22; lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1; lean_object* l_Lean_Attribute_Builtin_getIdent_x3f(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); @@ -3089,7 +3089,7 @@ x_6 = l_Lean_registerInitAttr(x_1, x_5, x_3, x_4); return x_6; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1() { _start: { lean_object* x_1; @@ -3097,17 +3097,17 @@ x_1 = lean_mk_string_from_bytes("init", 4); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3() { _start: { lean_object* x_1; @@ -3115,28 +3115,28 @@ x_1 = lean_mk_string_from_bytes("regularInitAttr", 15); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150_(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_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2; x_3 = 1; -x_4 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4; x_5 = l_Lean_registerInitAttrUnsafe(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1() { _start: { lean_object* x_1; @@ -3144,17 +3144,17 @@ x_1 = lean_mk_string_from_bytes("builtin_init", 12); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3() { _start: { lean_object* x_1; @@ -3162,23 +3162,23 @@ x_1 = lean_mk_string_from_bytes("builtinInitAttr", 15); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178_(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_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2; x_3 = 0; -x_4 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4; +x_4 = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4; x_5 = l_Lean_registerInitAttrUnsafe(x_2, x_3, x_4, x_1); return x_5; } @@ -3750,7 +3750,7 @@ lean_inc(x_56); x_57 = lean_ctor_get(x_55, 1); lean_inc(x_57); lean_dec(x_55); -x_58 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_56, x_3, x_4, x_57); +x_58 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_56, x_3, x_4, x_57); return x_58; } else @@ -3818,12 +3818,16 @@ lean_dec(x_3); return x_6; } } +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_InfoTree_Main(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_InitAttr(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_AddDecl(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Elab_InfoTree_Main(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -3930,28 +3934,28 @@ l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__29 = _init_l___auto lean_mark_persistent(l___auto____x40_Lean_Compiler_InitAttr___hyg_1136____closed__29); l___auto____x40_Lean_Compiler_InitAttr___hyg_1136_ = _init_l___auto____x40_Lean_Compiler_InitAttr___hyg_1136_(); lean_mark_persistent(l___auto____x40_Lean_Compiler_InitAttr___hyg_1136_); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__1); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__2); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__3); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180____closed__4); -if (builtin) {res = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1180_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__1); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__2); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__3); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150____closed__4); +if (builtin) {res = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1150_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_regularInitAttr = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_regularInitAttr); lean_dec_ref(res); -}l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__1); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__2); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__3); -l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208____closed__4); -if (builtin) {res = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1208_(lean_io_mk_world()); +}l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__1); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__2); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__3); +l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4 = _init_l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178____closed__4); +if (builtin) {res = l_Lean_initFn____x40_Lean_Compiler_InitAttr___hyg_1178_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_builtinInitAttr = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_builtinInitAttr); diff --git a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c index 29d104379816..42ed2238f59c 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/JoinPoints.c @@ -44,7 +44,6 @@ lean_object* l_Lean_Compiler_LCNF_Arg_forFVarM___at___private_Lean_Compiler_LCNF lean_object* l_Lean_Compiler_LCNF_ScopeM_withBackTrackingScope___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__3(lean_object*); lean_object* l_Lean_Compiler_LCNF_Arg_mapFVarM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__25; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1; LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___spec__2___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__2(lean_object*); @@ -94,7 +93,6 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Compiler_LCNF_JoinPointF lean_object* l_Lean_Compiler_LCNF_AltCore_mapCodeM___at_Lean_Compiler_LCNF_JoinPointFinder_replace_go___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__19; static lean_object* l_panic___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__6___closed__2; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1; static lean_object* l_Lean_Compiler_LCNF_extendJoinPointContext___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___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_array_fget(lean_object*, lean_object*); @@ -107,7 +105,6 @@ static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Compiler_LCNF_Decl_find LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointFinder_replace___spec__8___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_Decl_findJoinPoints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Compiler_LCNF_CompilerM_0__Lean_Compiler_LCNF_updateFunDeclImp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104_(lean_object*); lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___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*); static lean_object* l_Lean_Compiler_LCNF_Decl_findJoinPoints___closed__2; static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__22; @@ -139,6 +136,7 @@ lean_object* l_Lean_Compiler_LCNF_LCtx_toLocalContext(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_instInhabitedCandidateInfo; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__7(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__7___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___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_Compiler_LCNF_allFVarM___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -243,6 +241,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPoin lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_withNewJpScope___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_AltCore_getCode(lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2; lean_object* l_Lean_Compiler_LCNF_Decl_findJoinPoints___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashSet___at_Lean_Compiler_LCNF_JoinPointFinder_instInhabitedCandidateInfo___spec__1___boxed(lean_object*); lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,6 +281,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinP lean_object* l_Lean_Compiler_LCNF_FunDeclCore_getArity___rarg(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__9; static lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_find___closed__1; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__3(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_JoinPointFinder_FindState_scope___default; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyzeFunDecl___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -322,7 +322,6 @@ static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPo LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goReduce___spec__7(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInArg___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_AssocList_replace___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_addCandidate___spec__5(lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038_(lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__9___closed__1; static lean_object* l_panic___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___spec__4___closed__1; lean_object* l_Lean_Compiler_LCNF_LetValue_forFVarM___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInLetValue___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -349,7 +348,6 @@ uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Compiler_LCNF_JoinPointContextExtender_extendByIfNecessary___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2; LEAN_EXPORT lean_object* l_List_elem___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_addDependency___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extendByIfNecessary___spec__6(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4910____closed__6; @@ -413,12 +411,14 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Compiler_LCNF_JoinPoin lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Expr_fvar___override(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__24; +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991_(lean_object*); static lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Compiler_LCNF_JoinPointCommonArgs_reduce_goAnalyze___spec__14___boxed__const__1; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_Expr_mapFVarM___at_Lean_Compiler_LCNF_JoinPointContextExtender_extend_go___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_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addTrace___at_Lean_Compiler_LCNF_Decl_findJoinPoints___spec__2___closed__4; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_replaceFVar___closed__3; @@ -471,7 +471,6 @@ lean_object* l_Lean_Compiler_LCNF_ScopeM_withBackTrackingScope___at_Lean_Compile lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointFinder_find_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_Lean_Compiler_LCNF_findJoinPoints___closed__2; static lean_object* l_Lean_Compiler_LCNF_Decl_findJoinPoints___closed__1; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2; lean_object* l___private_Lean_Compiler_LCNF_Basic_0__Lean_Compiler_LCNF_Arg_updateFVarImp(lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compiler_LCNF_JoinPointFinder_removeCandidatesInArg___spec__3___closed__1; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4910____closed__3; @@ -480,6 +479,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_JoinPoints_0__Lean_Compi lean_object* l_Lean_Compiler_LCNF_JoinPointContextExtender_extendByIfNecessary(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4910____closed__2; +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057_(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4976____closed__7; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Compiler_LCNF_JoinPointCommonArgs_isInJpScope___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_ScopeM_withNewScope___at_Lean_Compiler_LCNF_JoinPointFinder_find_go___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -20652,7 +20652,7 @@ x_5 = l_Lean_Compiler_LCNF_extendJoinPointContext(x_1, x_4, x_3); return x_5; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20662,23 +20662,23 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4910____closed__16; -x_2 = lean_unsigned_to_nat(5038u); +x_2 = lean_unsigned_to_nat(4991u); 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_JoinPoints___hyg_5038_(lean_object* x_1) { +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -20737,7 +20737,7 @@ x_1 = l_Lean_Compiler_LCNF_commonJoinPointArgs___closed__4; return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -20747,23 +20747,23 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4910____closed__16; -x_2 = lean_unsigned_to_nat(5104u); +x_2 = lean_unsigned_to_nat(5057u); 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_JoinPoints___hyg_5104_(lean_object* x_1) { +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1; x_3 = 1; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -21025,11 +21025,11 @@ l_Lean_Compiler_LCNF_extendJoinPointContext___closed__2 = _init_l_Lean_Compiler_ lean_mark_persistent(l_Lean_Compiler_LCNF_extendJoinPointContext___closed__2); l_Lean_Compiler_LCNF_extendJoinPointContext___closed__3 = _init_l_Lean_Compiler_LCNF_extendJoinPointContext___closed__3(); lean_mark_persistent(l_Lean_Compiler_LCNF_extendJoinPointContext___closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038____closed__2); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5038_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991____closed__2); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_4991_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Compiler_LCNF_commonJoinPointArgs___closed__1 = _init_l_Lean_Compiler_LCNF_commonJoinPointArgs___closed__1(); @@ -21042,11 +21042,11 @@ l_Lean_Compiler_LCNF_commonJoinPointArgs___closed__4 = _init_l_Lean_Compiler_LCN lean_mark_persistent(l_Lean_Compiler_LCNF_commonJoinPointArgs___closed__4); l_Lean_Compiler_LCNF_commonJoinPointArgs = _init_l_Lean_Compiler_LCNF_commonJoinPointArgs(); lean_mark_persistent(l_Lean_Compiler_LCNF_commonJoinPointArgs); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104____closed__2); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5104_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057____closed__2); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_JoinPoints___hyg_5057_(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/Main.c b/stage0/stdlib/Lean/Compiler/LCNF/Main.c index 0ae7669e4e87..a652a7289d7d 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Main.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Main.c @@ -43,7 +43,6 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager static lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_PassManager_run___spec__8___closed__4; 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, 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*, 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*, lean_object*); 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*); @@ -66,12 +65,12 @@ extern lean_object* l_Lean_Expr_instBEq; extern lean_object* l_Lean_casesOnSuffix; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__20; lean_object* l_Lean_Compiler_LCNF_shouldGenerateCode(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(lean_object*, lean_object*, lean_object*, 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*); 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*); +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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; @@ -80,6 +79,7 @@ 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_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Main___hyg_1408____closed__1; +lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___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_Compiler_LCNF_checkpoint___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___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_lcnf_compile_decls(lean_object*, lean_object*, lean_object*, lean_object*); @@ -6170,7 +6170,7 @@ x_15 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___l x_16 = 1; x_17 = l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_checkpoint___spec__1___lambda__2___closed__6; x_18 = lean_box(x_16); -x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed), 8, 5); +x_19 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___boxed), 8, 5); lean_closure_set(x_19, 0, x_15); lean_closure_set(x_19, 1, x_6); lean_closure_set(x_19, 2, x_14); @@ -6178,7 +6178,7 @@ lean_closure_set(x_19, 3, x_18); lean_closure_set(x_19, 4, x_17); x_20 = l_Lean_Compiler_LCNF_main___closed__1; x_21 = lean_box(0); -x_22 = l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(x_20, x_5, x_19, x_21, x_2, x_3, x_4); +x_22 = l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(x_20, x_5, x_19, x_21, x_2, x_3, x_4); lean_dec(x_5); return x_22; } diff --git a/stage0/stdlib/Lean/Compiler/LCNF/Passes.c b/stage0/stdlib/Lean/Compiler/LCNF/Passes.c index 95a4c19d51e1..6b7d456b6780 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/Passes.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/Passes.c @@ -77,6 +77,7 @@ lean_object* l_Lean_Compiler_LCNF_Decl_saveMono(lean_object*, lean_object*, lean static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_498____closed__2; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_855____lambda__2___closed__2; extern lean_object* l_Lean_Compiler_LCNF_commonJoinPointArgs; +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_855____closed__12; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__15; @@ -164,7 +165,6 @@ LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_init; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__36; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__5; static lean_object* l_Lean_Compiler_LCNF_trace___closed__3; static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__17; @@ -216,10 +216,10 @@ static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__29; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_855____closed__8; lean_object* l_Lean_Compiler_LCNF_init___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_855____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_runImportedDecls___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_builtinPassManager___closed__25; static lean_object* l_Lean_Compiler_LCNF_trace___closed__1; +lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_1004____closed__2; static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_Passes___hyg_855____closed__1; lean_object* l_Lean_Compiler_LCNF_trace___boxed(lean_object*); @@ -1951,7 +1951,7 @@ x_23 = l_Lean_Compiler_LCNF_addPass___closed__4; x_24 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_24, 0, x_22); lean_ctor_set(x_24, 1, x_23); -x_25 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_24, x_2, x_3, x_7); +x_25 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_24, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_25; @@ -1977,7 +1977,7 @@ x_31 = l_Lean_Compiler_LCNF_addPass___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_ReducibilityAttrs_0__Lean_validate___spec__3(x_32, x_2, x_3, x_7); +x_33 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_32, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_33; @@ -2002,7 +2002,7 @@ x_39 = l_Lean_Compiler_LCNF_addPass___closed__4; x_40 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_40, 0, x_38); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_40, x_2, x_3, x_7); +x_41 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_40, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_41; @@ -2026,7 +2026,7 @@ x_47 = l_Lean_Compiler_LCNF_addPass___closed__4; x_48 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_48, x_2, x_3, x_7); +x_49 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_48, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_49; @@ -2157,7 +2157,7 @@ x_84 = l_Lean_Compiler_LCNF_addPass___closed__4; x_85 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_85, 0, x_83); lean_ctor_set(x_85, 1, x_84); -x_86 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_85, x_2, x_3, x_7); +x_86 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_85, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_86; @@ -2180,7 +2180,7 @@ x_90 = l_Lean_Compiler_LCNF_addPass___closed__4; x_91 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_91, 0, x_89); lean_ctor_set(x_91, 1, x_90); -x_92 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_91, x_2, x_3, x_7); +x_92 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_91, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_92; @@ -2202,7 +2202,7 @@ x_96 = l_Lean_Compiler_LCNF_addPass___closed__4; x_97 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_97, 0, x_95); lean_ctor_set(x_97, 1, x_96); -x_98 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_97, x_2, x_3, x_7); +x_98 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_97, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_98; @@ -2223,7 +2223,7 @@ x_102 = l_Lean_Compiler_LCNF_addPass___closed__4; x_103 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_103, 0, x_101); lean_ctor_set(x_103, 1, x_102); -x_104 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_103, x_2, x_3, x_7); +x_104 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_103, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_104; @@ -2243,7 +2243,7 @@ x_108 = l_Lean_Compiler_LCNF_addPass___closed__4; x_109 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_109, 0, x_107); lean_ctor_set(x_109, 1, x_108); -x_110 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_109, x_2, x_3, x_7); +x_110 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_109, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_110; @@ -2262,7 +2262,7 @@ x_114 = l_Lean_Compiler_LCNF_addPass___closed__4; x_115 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_115, 0, x_113); lean_ctor_set(x_115, 1, x_114); -x_116 = l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(x_115, x_2, x_3, x_7); +x_116 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_115, x_2, x_3, x_7); lean_dec(x_3); lean_dec(x_2); return x_116; @@ -2381,7 +2381,7 @@ x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); x_9 = 0; -x_10 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_9); +x_10 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; uint8_t x_13; diff --git a/stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c b/stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c index 68685aef478c..e639a5db4add 100644 --- a/stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c +++ b/stage0/stdlib/Lean/Compiler/LCNF/PhaseExt.c @@ -14,7 +14,6 @@ extern "C" { #endif static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__23; -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_forEachDecl___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_PhaseExt___hyg_172____closed__2; LEAN_EXPORT uint8_t l___private_Lean_Compiler_LCNF_PhaseExt_0__Lean_Compiler_LCNF_declLt(lean_object*, lean_object*); @@ -94,6 +93,7 @@ static lean_object* l_Lean_Compiler_LCNF_forEachModuleDecl___closed__4; static lean_object* l_panic___at_Lean_Compiler_LCNF_Decl_save___spec__1___closed__2; lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_forEachDecl___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1; lean_object* l_Lean_PersistentHashMap_instInhabited___rarg(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__2; lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -116,6 +116,7 @@ static lean_object* l_Lean_Compiler_LCNF_getBaseDecl_x3f___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_getExt___boxed(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Compiler_LCNF_getDeclCore_x3f___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at___private_Lean_Compiler_LCNF_PhaseExt_0__Lean_Compiler_LCNF_sortDecls___spec__1___closed__1; static lean_object* l_Lean_Compiler_LCNF_getDeclCore_x3f___closed__1; @@ -124,7 +125,6 @@ lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_ static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__22; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_mkDeclExt___spec__2___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_LCNF_forEachDecl___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_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__3; static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__28; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Compiler_LCNF_forEachModuleDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_mkDeclExt___spec__2___closed__2; @@ -158,15 +158,12 @@ lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); lean_object* l_Lean_Compiler_LCNF_Decl_save(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_LCNF_monoExt; -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Compiler_LCNF_forEachDecl___spec__5(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Compiler_LCNF_forEachDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2; lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__12; lean_object* lean_nat_sub(lean_object*, lean_object*); -static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at_Lean_Compiler_LCNF_forEachDecl___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_getExt___closed__2; @@ -175,6 +172,7 @@ extern lean_object* l_Lean_Core_instMonadCoreM; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_Compiler_LCNF_forEachDecl___spec__5___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Compiler_LCNF_mkDeclExt___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_PhaseExt_0__Lean_Compiler_LCNF_declLt___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Compiler_LCNF_PhaseExt_0__Lean_Compiler_LCNF_findAtSorted_x3f(lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); @@ -208,11 +206,13 @@ lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__25; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Compiler_LCNF_mkDeclExt___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4; lean_object* l_Lean_Compiler_LCNF_Decl_saveBase___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_LCNF_mkDeclExt___closed__1; lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_forEachModuleDecl(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__17; +static lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2; static lean_object* l_Lean_Compiler_LCNF_mkDeclExt___closed__7; static lean_object* l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__10; lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -1561,7 +1561,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1() { _start: { lean_object* x_1; @@ -1569,7 +1569,7 @@ x_1 = lean_mk_string_from_bytes("Compiler", 8); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2() { _start: { lean_object* x_1; @@ -1577,7 +1577,7 @@ 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_PhaseExt___hyg_172____closed__3() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3() { _start: { lean_object* x_1; @@ -1585,28 +1585,28 @@ x_1 = lean_mk_string_from_bytes("baseExt", 7); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1; -x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__3; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1; +x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172_(lean_object* x_1) { +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4; x_3 = l_Lean_Compiler_LCNF_mkDeclExt(x_2, x_1); return x_3; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1() { _start: { lean_object* x_1; @@ -1614,23 +1614,23 @@ x_1 = lean_mk_string_from_bytes("monoExt", 7); return x_1; } } -static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2() { +static lean_object* _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____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___auto____x40_Lean_Compiler_LCNF_PhaseExt___hyg_57____closed__1; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1; -x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2; -x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1; +x_3 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2; +x_4 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202_(lean_object* x_1) { +lean_object* l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2; +x_2 = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2; x_3 = l_Lean_Compiler_LCNF_mkDeclExt(x_2, x_1); return x_3; } @@ -3507,24 +3507,24 @@ l_Lean_Compiler_LCNF_mkDeclExt___closed__6 = _init_l_Lean_Compiler_LCNF_mkDeclEx lean_mark_persistent(l_Lean_Compiler_LCNF_mkDeclExt___closed__6); l_Lean_Compiler_LCNF_mkDeclExt___closed__7 = _init_l_Lean_Compiler_LCNF_mkDeclExt___closed__7(); lean_mark_persistent(l_Lean_Compiler_LCNF_mkDeclExt___closed__7); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__3(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__3); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__4); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172_(lean_io_mk_world()); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__2); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__3); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142____closed__4); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_142_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Compiler_LCNF_baseExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Compiler_LCNF_baseExt); lean_dec_ref(res); -}l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__1); -l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2(); -lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202____closed__2); -if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_202_(lean_io_mk_world()); +}l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__1); +l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2 = _init_l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2(); +lean_mark_persistent(l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172____closed__2); +if (builtin) {res = l_Lean_Compiler_LCNF_initFn____x40_Lean_Compiler_LCNF_PhaseExt___hyg_172_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Compiler_LCNF_monoExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Compiler_LCNF_monoExt); diff --git a/stage0/stdlib/Lean/Compiler/Main.c b/stage0/stdlib/Lean/Compiler/Main.c index 090dc76cbdab..12bde6a0ed91 100644 --- a/stage0/stdlib/Lean/Compiler/Main.c +++ b/stage0/stdlib/Lean/Compiler/Main.c @@ -14,13 +14,13 @@ extern "C" { #endif static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_89____closed__7; -lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_89____closed__6; lean_object* l_Lean_MessageData_ofList(lean_object*); static lean_object* l_Lean_Compiler_initFn____x40_Lean_Compiler_Main___hyg_89____closed__2; -lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__8___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_89____closed__3; lean_object* l_Lean_stringToMessageData(lean_object*); +lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_LCNF_compile(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Compiler_compile___lambda__1___closed__2; static lean_object* l_Lean_Compiler_compile___lambda__1___closed__1; @@ -205,7 +205,7 @@ x_8 = l_Lean_Compiler_compile___closed__2; x_9 = 1; x_10 = l_Lean_Compiler_compile___lambda__1___closed__3; x_11 = lean_box(x_9); -x_12 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed), 8, 5); +x_12 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__6___boxed), 8, 5); lean_closure_set(x_12, 0, x_8); lean_closure_set(x_12, 1, x_6); lean_closure_set(x_12, 2, x_7); @@ -213,7 +213,7 @@ lean_closure_set(x_12, 3, x_11); lean_closure_set(x_12, 4, x_10); x_13 = l_Lean_Compiler_compile___closed__3; x_14 = lean_box(0); -x_15 = l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(x_13, x_5, x_12, x_14, x_2, x_3, x_4); +x_15 = l_Lean_profileitM___at_Lean_addDecl___spec__8___rarg(x_13, x_5, x_12, x_14, x_2, x_3, x_4); lean_dec(x_5); return x_15; } diff --git a/stage0/stdlib/Lean/Compiler/Old.c b/stage0/stdlib/Lean/Compiler/Old.c index 92a5741c95f5..46bffdf887b9 100644 --- a/stage0/stdlib/Lean/Compiler/Old.c +++ b/stage0/stdlib/Lean/Compiler/Old.c @@ -28,9 +28,7 @@ static lean_object* l_Lean_Compiler_isEagerLambdaLiftingName___closed__1; lean_object* lean_compile_decls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_compileDecls___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, lean_object*); static lean_object* l_Lean_Compiler_checkIsDefinition___closed__4; -LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_is_unsafe_rec_name(lean_object*); static lean_object* l_Lean_Compiler_checkIsDefinition___closed__1; static lean_object* l_Lean_Compiler_checkIsDefinition___closed__2; @@ -38,7 +36,6 @@ uint8_t l_String_isPrefixOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Compiler_checkIsDefinition(lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_mk_unsafe_rec_name(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT uint8_t lean_is_eager_lambda_lifting_name(lean_object*); static lean_object* l_Lean_Compiler_checkIsDefinition___closed__3; @@ -417,51 +414,6 @@ lean_dec(x_2); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = lean_add_decl(x_1, x_3); -if (lean_obj_tag(x_4) == 0) -{ -uint8_t x_5; -lean_dec(x_3); -x_5 = !lean_is_exclusive(x_4); -if (x_5 == 0) -{ -return x_4; -} -else -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_4, 0); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_7, 0, x_6); -return x_7; -} -} -else -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_4, 0); -lean_inc(x_8); -lean_dec(x_4); -x_9 = l_Lean_Environment_compileDecl(x_8, x_2, x_3); -return x_9; -} -} -} -LEAN_EXPORT lean_object* l_Lean_Environment_addAndCompile___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Environment_addAndCompile(x_1, x_2, x_3); -lean_dec(x_2); -return x_4; -} -} lean_object* initialize_Lean_Environment(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Compiler_Old(uint8_t builtin, lean_object* w) { diff --git a/stage0/stdlib/Lean/CoreM.c b/stage0/stdlib/Lean/CoreM.c index 1ae70d9c91de..ef3cf8fea548 100644 --- a/stage0/stdlib/Lean/CoreM.c +++ b/stage0/stdlib/Lean/CoreM.c @@ -17,10 +17,8 @@ LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___boxed(lean_object*); lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3; -static lean_object* l_Lean_addDecl___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_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; @@ -30,15 +28,12 @@ LEAN_EXPORT uint8_t l_Lean_Exception_isMaxHeartbeat(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Core_instantiateTypeLevelParams___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_ref___default; static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__4; -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5; static lean_object* l_Lean_instMonadExceptOfExceptionCoreM___closed__1; -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCoreM(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currMacroScope___default; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Core_instantiateTypeLevelParams___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__5; LEAN_EXPORT lean_object* l_Lean_instMonadExceptOfExceptionCoreM; extern lean_object* l_Lean_profiler; @@ -51,12 +46,13 @@ LEAN_EXPORT lean_object* l_Lean_Core_Context_options___default; static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2___closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__3; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_right(size_t, size_t); LEAN_EXPORT lean_object* l_List_elem___at_Lean_catchInternalIds___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__5(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__10; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__8; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694_(lean_object*); static lean_object* l_Lean_ImportM_runCoreM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog(lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__14; @@ -65,7 +61,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); -lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Exception_isMaxRecDepth(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM(lean_object*); lean_object* lean_io_check_canceled(lean_object*); @@ -78,22 +73,18 @@ extern lean_object* l_Lean_maxRecDepthErrorMessage; LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2; lean_object* l_Lean_MessageData_toString(lean_object*, lean_object*); double lean_float_div(double, double); -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_checkSystem___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -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_Lean_Core_instInhabitedCoreM___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__3; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Cache_instLevelType___default; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); @@ -101,7 +92,6 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__9; LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth(lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadInfoTreeCoreM___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM; lean_object* l_Lean_MessageData_ofList(lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); @@ -119,38 +109,33 @@ LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___boxed(lean_obj static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecl(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_casesOnSuffix; -lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM; static lean_object* l_Lean_instMonadExceptOfExceptionCoreM___closed__3; lean_object* l_Lean_ConstantInfo_name(lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadEnvCoreM___closed__2; -static double l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6_(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*); LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); static lean_object* l_Lean_Core_instAddMessageContextCoreM___closed__1; uint8_t l_Lean_ConstantInfo_hasValue(lean_object*); static lean_object* l_Lean_Core_instMonadInfoTreeCoreM___closed__3; @@ -161,7 +146,6 @@ static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__4; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withCurrHeartbeatsImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Compiler_getDeclNamesForCodeGen(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__3(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addDecl___closed__2; lean_object* lean_io_get_num_heartbeats(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; lean_object* lean_lcnf_compile_decls(lean_object*, lean_object*, lean_object*, lean_object*); @@ -169,7 +153,6 @@ lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__3(lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); lean_object* l_EStateM_instMonad(lean_object*, lean_object*); -static lean_object* l_Lean_addDecl___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_mkFreshUserName___boxed(lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepth; @@ -179,7 +162,6 @@ static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__ LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Core_instantiateTypeLevelParams___spec__3(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Core_instantiateTypeLevelParams___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Core_instantiateTypeLevelParams___spec__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_checkInterrupted___rarg___closed__3; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14; @@ -190,7 +172,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Core_instantiateTypeLevelParams___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRefCoreM___closed__1; lean_object* l_Lean_Environment_compileDecl(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); @@ -211,31 +192,32 @@ static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__6; static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__1; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withConsistentCtx___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__5; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_traceState___default; +LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(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*); +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3; LEAN_EXPORT lean_object* l_Lean_mapCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Core_Context_catchRuntimeEx___default; -LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_modifyInstLevelTypeCache___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_tryCatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__6; static lean_object* l_Lean_Core_checkInterrupted___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withCatchingRuntimeEx___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___closed__2; 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_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*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___closed__4; static lean_object* l_Lean_Core_instMonadQuotationCoreM___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_beq___at_Lean_Core_instantiateTypeLevelParams___spec__8(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalId___boxed(lean_object*, lean_object*, lean_object*); @@ -252,28 +234,22 @@ uint8_t lean_is_aux_recursor(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM; -LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__5; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29; -static lean_object* l_Lean_addDecl___closed__1; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__7; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2; -static double l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_Context_openDecls___default; -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4; LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__3___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_compileDecl___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_compile_decls(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currRecDepth___default; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898_(lean_object*); static lean_object* l_Lean_Core_instInhabitedCoreM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Core_instantiateTypeLevelParams___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -282,14 +258,14 @@ lean_object* l_Lean_PersistentArray_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Core_State_ngen___default___closed__1; LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__4; +LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20; static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_41____closed__2; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_forallE___override(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__2; +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_compileDecl___lambda__1___closed__2; static lean_object* l_Lean_Core_instMonadRefCoreM___closed__3; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___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_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -303,7 +279,6 @@ static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__13 lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21; LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__1; LEAN_EXPORT lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_cache___default; @@ -317,7 +292,6 @@ LEAN_EXPORT lean_object* l_Lean_printTraces___at_Lean_Core_instMetaEvalCoreM___s LEAN_EXPORT lean_object* l_Lean_Core_instantiateTypeLevelParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_41____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___closed__1; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Core_instantiateTypeLevelParams___spec__6(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_checkInterrupted(lean_object*, lean_object*); @@ -325,23 +299,22 @@ lean_object* lean_io_mono_nanos_now(lean_object*); static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__3; static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__1; LEAN_EXPORT lean_object* l_Lean_mkArrowN(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_Cache_instLevelType___default___closed__2; static lean_object* l_Lean_ImportM_runCoreM___rarg___closed__1; static lean_object* l_Lean_Core_instMonadLogCoreM___closed__3; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run_x27(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_State_infoState___default; static lean_object* l_Lean_Core_instMonadInfoTreeCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__6; LEAN_EXPORT uint8_t l_Lean_Core_Context_diag___default; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRecDepthCoreM___closed__2; +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4; static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; LEAN_EXPORT lean_object* l_Lean_Core_withFreshMacroScope___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__1; +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Exception_isMaxHeartbeat___boxed(lean_object*); static lean_object* l_Lean_Core_State_ngen___default___closed__2; @@ -351,8 +324,6 @@ static uint8_t l_Lean_ImportM_runCoreM___rarg___closed__3; static lean_object* l_Lean_Core_checkInterrupted___rarg___closed__4; static lean_object* l_Lean_instMonadExceptOfExceptionCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_getDiag___boxed(lean_object*); -extern lean_object* l_Lean_warningAsError; -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM; LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_maxHeartbeats___default___boxed(lean_object*); @@ -372,37 +343,33 @@ LEAN_EXPORT lean_object* l_Lean_catchInternalId___rarg___lambda__1(lean_object*, static lean_object* l_Lean_mkArrow___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_setMessageLog(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalId___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_compileDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Exception_isMaxHeartbeat___closed__1; LEAN_EXPORT lean_object* l_Lean_diagnostics; 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*); static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___closed__3; -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__8___boxed(lean_object*, lean_object*, 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___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecls(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withCatchingRuntimeEx___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; -static lean_object* l_Lean_addDecl___lambda__1___closed__1; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_toIO(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_maxHeartbeats___default(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRecDepthCoreM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Core_instantiateValueLevelParams___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withAtLeastMaxRecDepth___rarg(lean_object*, lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_mkFreshUserName(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5; LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_inheritedTraceOptions; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Core_instantiateTypeLevelParams___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadWithOptionsCoreM___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_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*); @@ -417,6 +384,7 @@ LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_Core_withConsistentCtx(l static lean_object* l_Lean_Core_instMonadRecDepthCoreM___closed__1; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_80_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_getAndEmptyMessageLog___boxed(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28; lean_object* l_IO_println___at_Lean_instEvalOfRepr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDeclsNew___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -431,20 +399,19 @@ static lean_object* l_Lean_mkArrow___closed__2; LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_State_infoState___default___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2; +static lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2; static lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instantiateValueLevelParams___lambda__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalIds___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___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_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_CoreM_run___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadLogCoreM___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withAtLeastMaxRecDepth___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadRefCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog___boxed(lean_object*, lean_object*, lean_object*); @@ -461,14 +428,13 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCor uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__2(lean_object*, lean_object*, lean_object*); -static double l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5; +static double l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_State_traceState___default___closed__3; lean_object* l_ReaderT_instFunctorOfMonad___rarg(lean_object*); lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadLiftIOCoreM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_addDecl___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_catchInternalIds___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadQuotationCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -478,11 +444,12 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadLogCoreM___closed__5; static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1; LEAN_EXPORT lean_object* l_Lean_Core_checkMaxHeartbeatsCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_compileDecl___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_String_isPrefixOf(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isDiagnosticsEnabled___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Name_hash___override(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -493,6 +460,7 @@ static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_41_(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18; LEAN_EXPORT lean_object* l_Lean_Core_State_nextMacroScope___default; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_modifyCache(lean_object*, lean_object*, lean_object*, lean_object*); @@ -504,23 +472,22 @@ 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; LEAN_EXPORT lean_object* l_Lean_Core_resetMessageLog(lean_object*, lean_object*, lean_object*); +static double l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Core_Context_maxRecDepth___default; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___closed__1; +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_restore___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadCoreM___closed__3; static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__1; LEAN_EXPORT lean_object* l_Lean_mkArrow(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Core_instMetaEvalCoreM___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM; static lean_object* l_Lean_Core_checkInterrupted___rarg___closed__2; -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3; +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3; static lean_object* l_Lean_Core_instMonadCoreM___closed__6; -uint8_t l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(lean_object*, uint8_t); static lean_object* l_Lean_Core_State_traceState___default___closed__2; LEAN_EXPORT lean_object* l_Lean_instMonadExceptOfExceptionCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_isRecCore(lean_object*, lean_object*); @@ -528,34 +495,29 @@ lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_throwMaxHeartbeat(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_firstFrontendMacroScope; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Core_instMonadInfoTreeCoreM; LEAN_EXPORT lean_object* l_Lean_diagnostics_threshold; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadWithOptionsCoreM___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Exception_isRuntime___boxed(lean_object*); static lean_object* l_Lean_Core_instMonadLogCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_CoreM_toIO___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM___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_Core_withFreshMacroScope(lean_object*); size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at_Lean_mkArrowN___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadWithOptionsCoreM(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mapCoreM(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__1; LEAN_EXPORT lean_object* l_Lean_Core_Context_initHeartbeats___default; LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*); @@ -565,16 +527,17 @@ LEAN_EXPORT lean_object* l_Lean_Core_instMonadTraceCoreM___lambda__2(lean_object static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instMonadNameGeneratorCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__12; lean_object* lean_io_error_to_string(lean_object*); +static double l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadLogCoreM___closed__1; LEAN_EXPORT lean_object* l_Lean_Core_instMetaEvalCoreM___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_liftIOCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Core_instMetaEvalCoreM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Core_instantiateValueLevelParams___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__2(lean_object*, lean_object*, lean_object*); @@ -585,7 +548,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Core_instan LEAN_EXPORT lean_object* l_Lean_Core_instMonadLogCoreM___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Core_instMetaEvalCoreM___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors; -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1; static lean_object* l_Lean_Core_State_traceState___default___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__3(lean_object*); @@ -599,13 +561,11 @@ LEAN_EXPORT lean_object* l_Lean_Core_instAddMessageContextCoreM; LEAN_EXPORT lean_object* l_Lean_instMonadExceptOfExceptionCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadEnvCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8; -static lean_object* l_Lean_addDecl___lambda__3___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2; -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withCurrHeartbeats___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportM_runCoreM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113_(lean_object*); -uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); uint8_t lean_level_eq(lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__4; @@ -614,29 +574,28 @@ static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__2; LEAN_EXPORT lean_object* l_Lean_Core_instInhabitedCache; 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*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withoutCatchingRuntimeEx(lean_object*, lean_object*); static lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__2; lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); -static lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_compileDecl___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23; LEAN_EXPORT lean_object* l_Lean_Core_instMonadCoreM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadQuotationCoreM___closed__4; LEAN_EXPORT uint8_t l_Lean_Exception_isRuntime(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__11; static lean_object* l_Lean_Core_instInhabitedCoreM___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_withAtLeastMaxRecDepth(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isDiagnosticsEnabled(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_getMessageLog___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_restoreFull(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__6; +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_compileDecl___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadRefCoreM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_maxHeartbeats; static lean_object* l_Lean_Core_instMonadResolveNameCoreM___closed__3; -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_compileDecl___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_compileDecl___lambda__3___closed__1; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1___closed__4; lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -647,19 +606,16 @@ lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Lean_Core_throwMaxHeartbeat___closed__5; LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2(lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27; -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Core_instMonadWithOptionsCoreM___rarg___closed__1; +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Core_withIncRecDepth___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1; -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_instMonadResolveNameCoreM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__13; lean_object* l_Lean_MessageData_format(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ReaderT_instMonad___rarg(lean_object*); size_t lean_usize_land(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_compiler_enableNew; LEAN_EXPORT lean_object* l_Lean_Core_getMaxHeartbeats___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_Context_currNamespace___default; @@ -11242,3378 +11198,540 @@ lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1() { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 5); -x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_inc(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); -lean_ctor_set_tag(x_6, 1); -lean_ctor_set(x_6, 0, x_9); -return x_6; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Empty", 5); +return x_1; } -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_6, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_6); -lean_inc(x_5); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; } +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("rec", 3); +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_2, 2); -lean_inc(x_5); -x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); -x_7 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_6, x_2, x_3, x_4); -lean_dec(x_2); -return x_7; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4() { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_5 = lean_st_ref_take(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = !lean_is_exclusive(x_6); -if (x_8 == 0) -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; -x_9 = lean_ctor_get(x_6, 4); -lean_dec(x_9); -x_10 = lean_ctor_get(x_6, 0); -lean_dec(x_10); -x_11 = l_Lean_Core_instInhabitedCache___closed__1; -lean_ctor_set(x_6, 4, x_11); -lean_ctor_set(x_6, 0, x_1); -x_12 = lean_st_ref_set(x_3, x_6, x_7); -x_13 = !lean_is_exclusive(x_12); -if (x_13 == 0) +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("False", 5); +return x_1; +} +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__5() { +_start: { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_12, 0); -lean_dec(x_14); -x_15 = lean_box(0); -lean_ctor_set(x_12, 0, x_15); -return x_12; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6() { +_start: { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_12, 1); -lean_inc(x_16); -lean_dec(x_12); -x_17 = lean_box(0); -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; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Eq", 2); +return x_1; } } -else +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__7() { +_start: { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_19 = lean_ctor_get(x_6, 1); -x_20 = lean_ctor_get(x_6, 2); -x_21 = lean_ctor_get(x_6, 3); -x_22 = lean_ctor_get(x_6, 5); -x_23 = lean_ctor_get(x_6, 6); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_6); -x_24 = l_Lean_Core_instInhabitedCache___closed__1; -x_25 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_25, 0, x_1); -lean_ctor_set(x_25, 1, x_19); -lean_ctor_set(x_25, 2, x_20); -lean_ctor_set(x_25, 3, x_21); -lean_ctor_set(x_25, 4, x_24); -lean_ctor_set(x_25, 5, x_22); -lean_ctor_set(x_25, 6, x_23); -x_26 = lean_st_ref_set(x_3, x_25, x_7); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_28 = x_26; -} else { - lean_dec_ref(x_26); - x_28 = lean_box(0); +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("ndrec", 5); +return x_1; } -x_29 = lean_box(0); -if (lean_is_scalar(x_28)) { - x_30 = lean_alloc_ctor(0, 2, 0); -} else { - x_30 = x_28; } -lean_ctor_set(x_30, 0, x_29); -lean_ctor_set(x_30, 1, x_27); -return x_30; +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__7; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} } -static lean_object* _init_l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1() { +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__10() { _start: { lean_object* x_1; -x_1 = l_Lean_warningAsError; +x_1 = lean_mk_string_from_bytes("recOn", 5); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__11() { _start: { -lean_object* x_7; uint8_t x_265; uint8_t x_266; -x_265 = 2; -x_266 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_265); -if (x_266 == 0) -{ -lean_object* x_267; -x_267 = lean_box(0); -x_7 = x_267; -goto block_264; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__10; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else -{ -lean_object* x_268; uint8_t x_269; -lean_inc(x_2); -x_268 = l_Lean_MessageData_hasSyntheticSorry(x_2); -x_269 = lean_unbox(x_268); -lean_dec(x_268); -if (x_269 == 0) -{ -lean_object* x_270; -x_270 = lean_box(0); -x_7 = x_270; -goto block_264; } -else +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12() { +_start: { -lean_object* x_271; lean_object* x_272; -lean_dec(x_2); -x_271 = lean_box(0); -x_272 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_272, 0, x_271); -lean_ctor_set(x_272, 1, x_6); -return x_272; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("casesOn", 7); +return x_1; } } -block_264: +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__13() { +_start: { -uint8_t x_8; lean_object* x_258; uint8_t x_259; uint8_t x_260; -lean_dec(x_7); -x_258 = lean_ctor_get(x_4, 2); -x_259 = 1; -x_260 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_259); -if (x_260 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14() { +_start: { -x_8 = x_3; -goto block_257; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15() { +_start: { -lean_object* x_261; uint8_t x_262; -x_261 = l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; -x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); -if (x_262 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16() { +_start: { -x_8 = x_3; -goto block_257; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("And", 3); +return x_1; } -else +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__17() { +_start: { -uint8_t x_263; -x_263 = 2; -x_8 = x_263; -goto block_257; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; } } -block_257: +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_9 = lean_ctor_get(x_4, 0); -x_10 = lean_ctor_get(x_4, 1); -x_11 = lean_ctor_get(x_4, 5); -x_12 = lean_ctor_get(x_4, 6); -x_13 = lean_ctor_get(x_4, 7); -x_14 = l_Lean_replaceRef(x_1, x_11); -x_15 = 0; -x_16 = l_Lean_Syntax_getPos_x3f(x_14, x_15); -x_17 = l_Lean_Syntax_getTailPos_x3f(x_14, x_15); -if (lean_obj_tag(x_16) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__19() { +_start: { -if (lean_obj_tag(x_17) == 0) +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(10u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20() { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_18 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -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_unsigned_to_nat(0u); -x_22 = l_Lean_FileMap_toPosition(x_10, x_21); -lean_inc(x_22); -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_22); -lean_inc(x_13); -lean_inc(x_12); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_12); -lean_ctor_set(x_24, 1, x_13); -x_25 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_25, 0, x_24); -lean_ctor_set(x_25, 1, x_19); -x_26 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_27 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_27, 0, x_9); -lean_ctor_set(x_27, 1, x_22); -lean_ctor_set(x_27, 2, x_23); -lean_ctor_set(x_27, 3, x_26); -lean_ctor_set(x_27, 4, x_25); -lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_8); -x_28 = lean_st_ref_take(x_5, x_20); -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; uint8_t x_35; -x_32 = lean_ctor_get(x_29, 5); -x_33 = l_Lean_PersistentArray_push___rarg(x_32, x_27); -lean_ctor_set(x_29, 5, x_33); -x_34 = lean_st_ref_set(x_5, x_29, x_30); -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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__19; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21() { +_start: { -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; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__5; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__22() { +_start: { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; 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_29, 0); -x_42 = lean_ctor_get(x_29, 1); -x_43 = lean_ctor_get(x_29, 2); -x_44 = lean_ctor_get(x_29, 3); -x_45 = lean_ctor_get(x_29, 4); -x_46 = lean_ctor_get(x_29, 5); -x_47 = lean_ctor_get(x_29, 6); -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_inc(x_41); -lean_dec(x_29); -x_48 = l_Lean_PersistentArray_push___rarg(x_46, x_27); -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_44); -lean_ctor_set(x_49, 4, x_45); -lean_ctor_set(x_49, 5, x_48); -lean_ctor_set(x_49, 6, x_47); -x_50 = lean_st_ref_set(x_5, x_49, x_30); -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); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -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; +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__22; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__9; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -else -{ -uint8_t x_55; -x_55 = !lean_is_exclusive(x_17); -if (x_55 == 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; lean_object* x_69; uint8_t x_70; -x_56 = lean_ctor_get(x_17, 0); -x_57 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); -lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_unsigned_to_nat(0u); -x_61 = l_Lean_FileMap_toPosition(x_10, x_60); -x_62 = l_Lean_FileMap_toPosition(x_10, x_56); -lean_dec(x_56); -lean_ctor_set(x_17, 0, x_62); -lean_inc(x_13); -lean_inc(x_12); -x_63 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_63, 0, x_12); -lean_ctor_set(x_63, 1, x_13); -x_64 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_64, 0, x_63); -lean_ctor_set(x_64, 1, x_58); -x_65 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_66 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_66, 0, x_9); -lean_ctor_set(x_66, 1, x_61); -lean_ctor_set(x_66, 2, x_17); -lean_ctor_set(x_66, 3, x_65); -lean_ctor_set(x_66, 4, x_64); -lean_ctor_set_uint8(x_66, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_66, sizeof(void*)*5 + 1, x_8); -x_67 = lean_st_ref_take(x_5, x_59); -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_is_exclusive(x_68); -if (x_70 == 0) -{ -lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; -x_71 = lean_ctor_get(x_68, 5); -x_72 = l_Lean_PersistentArray_push___rarg(x_71, x_66); -lean_ctor_set(x_68, 5, x_72); -x_73 = lean_st_ref_set(x_5, x_68, x_69); -x_74 = !lean_is_exclusive(x_73); -if (x_74 == 0) +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__24() { +_start: { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_73, 0); -lean_dec(x_75); -x_76 = lean_box(0); -lean_ctor_set(x_73, 0, x_76); -return x_73; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__11; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -else +} +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__25() { +_start: { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_73, 1); -lean_inc(x_77); -lean_dec(x_73); -x_78 = lean_box(0); -x_79 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_79, 0, x_78); -lean_ctor_set(x_79, 1, x_77); -return x_79; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__24; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__13; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26() { +_start: { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_80 = lean_ctor_get(x_68, 0); -x_81 = lean_ctor_get(x_68, 1); -x_82 = lean_ctor_get(x_68, 2); -x_83 = lean_ctor_get(x_68, 3); -x_84 = lean_ctor_get(x_68, 4); -x_85 = lean_ctor_get(x_68, 5); -x_86 = lean_ctor_get(x_68, 6); -lean_inc(x_86); -lean_inc(x_85); -lean_inc(x_84); -lean_inc(x_83); -lean_inc(x_82); -lean_inc(x_81); -lean_inc(x_80); -lean_dec(x_68); -x_87 = l_Lean_PersistentArray_push___rarg(x_85, x_66); -x_88 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_88, 0, x_80); -lean_ctor_set(x_88, 1, x_81); -lean_ctor_set(x_88, 2, x_82); -lean_ctor_set(x_88, 3, x_83); -lean_ctor_set(x_88, 4, x_84); -lean_ctor_set(x_88, 5, x_87); -lean_ctor_set(x_88, 6, x_86); -x_89 = lean_st_ref_set(x_5, x_88, x_69); -x_90 = lean_ctor_get(x_89, 1); -lean_inc(x_90); -if (lean_is_exclusive(x_89)) { - lean_ctor_release(x_89, 0); - lean_ctor_release(x_89, 1); - x_91 = x_89; -} else { - lean_dec_ref(x_89); - x_91 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__25; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -x_92 = lean_box(0); -if (lean_is_scalar(x_91)) { - x_93 = lean_alloc_ctor(0, 2, 0); -} else { - x_93 = x_91; } -lean_ctor_set(x_93, 0, x_92); -lean_ctor_set(x_93, 1, x_90); -return x_93; +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15; +x_3 = lean_array_push(x_1, x_2); +return x_3; } } -else +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28() { +_start: { -lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_94 = lean_ctor_get(x_17, 0); -lean_inc(x_94); -lean_dec(x_17); -x_95 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -lean_dec(x_95); -x_98 = lean_unsigned_to_nat(0u); -x_99 = l_Lean_FileMap_toPosition(x_10, x_98); -x_100 = l_Lean_FileMap_toPosition(x_10, x_94); -lean_dec(x_94); -x_101 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_101, 0, x_100); -lean_inc(x_13); -lean_inc(x_12); -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_12); -lean_ctor_set(x_102, 1, x_13); -x_103 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_103, 0, x_102); -lean_ctor_set(x_103, 1, x_96); -x_104 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_105 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_105, 0, x_9); -lean_ctor_set(x_105, 1, x_99); -lean_ctor_set(x_105, 2, x_101); -lean_ctor_set(x_105, 3, x_104); -lean_ctor_set(x_105, 4, x_103); -lean_ctor_set_uint8(x_105, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_105, sizeof(void*)*5 + 1, x_8); -x_106 = lean_st_ref_take(x_5, x_97); -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_ctor_get(x_107, 0); -lean_inc(x_109); -x_110 = lean_ctor_get(x_107, 1); -lean_inc(x_110); -x_111 = lean_ctor_get(x_107, 2); -lean_inc(x_111); -x_112 = lean_ctor_get(x_107, 3); -lean_inc(x_112); -x_113 = lean_ctor_get(x_107, 4); -lean_inc(x_113); -x_114 = lean_ctor_get(x_107, 5); -lean_inc(x_114); -x_115 = lean_ctor_get(x_107, 6); -lean_inc(x_115); -if (lean_is_exclusive(x_107)) { - lean_ctor_release(x_107, 0); - lean_ctor_release(x_107, 1); - lean_ctor_release(x_107, 2); - lean_ctor_release(x_107, 3); - lean_ctor_release(x_107, 4); - lean_ctor_release(x_107, 5); - lean_ctor_release(x_107, 6); - x_116 = x_107; -} else { - lean_dec_ref(x_107); - x_116 = lean_box(0); +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__17; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -x_117 = l_Lean_PersistentArray_push___rarg(x_114, x_105); -if (lean_is_scalar(x_116)) { - x_118 = lean_alloc_ctor(0, 7, 0); -} else { - x_118 = x_116; -} -lean_ctor_set(x_118, 0, x_109); -lean_ctor_set(x_118, 1, x_110); -lean_ctor_set(x_118, 2, x_111); -lean_ctor_set(x_118, 3, x_112); -lean_ctor_set(x_118, 4, x_113); -lean_ctor_set(x_118, 5, x_117); -lean_ctor_set(x_118, 6, x_115); -x_119 = lean_st_ref_set(x_5, x_118, x_108); -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; +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28; +x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18; +x_3 = lean_array_push(x_1, x_2); +return x_3; } -lean_ctor_set(x_123, 0, x_122); -lean_ctor_set(x_123, 1, x_120); -return x_123; } +static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors() { +_start: +{ +lean_object* x_1; +x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29; +return x_1; } } -else +LEAN_EXPORT uint8_t l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { +_start: { -if (lean_obj_tag(x_17) == 0) +if (lean_obj_tag(x_2) == 4) { -uint8_t x_124; -x_124 = !lean_is_exclusive(x_16); -if (x_124 == 0) +lean_object* x_3; uint8_t x_4; +x_3 = lean_ctor_get(x_2, 0); +lean_inc(x_3); +lean_dec(x_2); +lean_inc(x_3); +lean_inc(x_1); +x_4 = lean_is_aux_recursor(x_1, x_3); +if (x_4 == 0) { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; -x_125 = lean_ctor_get(x_16, 0); -x_126 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_126, 1); -lean_inc(x_128); -lean_dec(x_126); -x_129 = l_Lean_FileMap_toPosition(x_10, x_125); -lean_dec(x_125); -lean_inc(x_129); -lean_ctor_set(x_16, 0, x_129); -lean_inc(x_13); -lean_inc(x_12); -x_130 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_130, 0, x_12); -lean_ctor_set(x_130, 1, x_13); -x_131 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_127); -x_132 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_133 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_133, 0, x_9); -lean_ctor_set(x_133, 1, x_129); -lean_ctor_set(x_133, 2, x_16); -lean_ctor_set(x_133, 3, x_132); -lean_ctor_set(x_133, 4, x_131); -lean_ctor_set_uint8(x_133, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_133, sizeof(void*)*5 + 1, x_8); -x_134 = lean_st_ref_take(x_5, x_128); -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_ctor_get(x_134, 1); -lean_inc(x_136); -lean_dec(x_134); -x_137 = !lean_is_exclusive(x_135); -if (x_137 == 0) +uint8_t x_5; +lean_inc(x_3); +x_5 = l_Lean_isRecCore(x_1, x_3); +if (x_5 == 0) +{ +uint8_t x_6; +lean_dec(x_3); +x_6 = 0; +return x_6; +} +else { -lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; -x_138 = lean_ctor_get(x_135, 5); -x_139 = l_Lean_PersistentArray_push___rarg(x_138, x_133); -lean_ctor_set(x_135, 5, x_139); -x_140 = lean_st_ref_set(x_5, x_135, x_136); -x_141 = !lean_is_exclusive(x_140); -if (x_141 == 0) +lean_object* x_7; uint8_t x_8; +x_7 = l___private_Lean_CoreM_0__Lean_supportedRecursors; +x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_3); +lean_dec(x_3); +if (x_8 == 0) { -lean_object* x_142; lean_object* x_143; -x_142 = lean_ctor_get(x_140, 0); -lean_dec(x_142); -x_143 = lean_box(0); -lean_ctor_set(x_140, 0, x_143); -return x_140; +uint8_t x_9; +x_9 = 1; +return x_9; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_140, 1); -lean_inc(x_144); -lean_dec(x_140); -x_145 = lean_box(0); -x_146 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_146, 0, x_145); -lean_ctor_set(x_146, 1, x_144); -return x_146; +uint8_t x_10; +x_10 = 0; +return x_10; +} } } else { -lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; -x_147 = lean_ctor_get(x_135, 0); -x_148 = lean_ctor_get(x_135, 1); -x_149 = lean_ctor_get(x_135, 2); -x_150 = lean_ctor_get(x_135, 3); -x_151 = lean_ctor_get(x_135, 4); -x_152 = lean_ctor_get(x_135, 5); -x_153 = lean_ctor_get(x_135, 6); -lean_inc(x_153); -lean_inc(x_152); -lean_inc(x_151); -lean_inc(x_150); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_135); -x_154 = l_Lean_PersistentArray_push___rarg(x_152, x_133); -x_155 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_155, 0, x_147); -lean_ctor_set(x_155, 1, x_148); -lean_ctor_set(x_155, 2, x_149); -lean_ctor_set(x_155, 3, x_150); -lean_ctor_set(x_155, 4, x_151); -lean_ctor_set(x_155, 5, x_154); -lean_ctor_set(x_155, 6, x_153); -x_156 = lean_st_ref_set(x_5, x_155, x_136); -x_157 = lean_ctor_get(x_156, 1); -lean_inc(x_157); -if (lean_is_exclusive(x_156)) { - lean_ctor_release(x_156, 0); - lean_ctor_release(x_156, 1); - x_158 = x_156; -} else { - lean_dec_ref(x_156); - x_158 = lean_box(0); -} -x_159 = lean_box(0); -if (lean_is_scalar(x_158)) { - x_160 = lean_alloc_ctor(0, 2, 0); -} else { - x_160 = x_158; -} -lean_ctor_set(x_160, 0, x_159); -lean_ctor_set(x_160, 1, x_157); -return x_160; -} +lean_object* x_11; uint8_t x_12; +x_11 = l_Lean_casesOnSuffix; +lean_inc(x_3); +lean_inc(x_1); +x_12 = l_Lean_isAuxRecursorWithSuffix(x_1, x_3, x_11); +if (x_12 == 0) +{ +lean_object* x_13; uint8_t x_14; +lean_dec(x_1); +x_13 = l___private_Lean_CoreM_0__Lean_supportedRecursors; +x_14 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_13, x_3); +lean_dec(x_3); +if (x_14 == 0) +{ +uint8_t x_15; +x_15 = 1; +return x_15; } else { -lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; -x_161 = lean_ctor_get(x_16, 0); -lean_inc(x_161); -lean_dec(x_16); -x_162 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -x_163 = lean_ctor_get(x_162, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_162, 1); -lean_inc(x_164); -lean_dec(x_162); -x_165 = l_Lean_FileMap_toPosition(x_10, x_161); -lean_dec(x_161); -lean_inc(x_165); -x_166 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_166, 0, x_165); -lean_inc(x_13); -lean_inc(x_12); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_12); -lean_ctor_set(x_167, 1, x_13); -x_168 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_168, 0, x_167); -lean_ctor_set(x_168, 1, x_163); -x_169 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_170 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_170, 0, x_9); -lean_ctor_set(x_170, 1, x_165); -lean_ctor_set(x_170, 2, x_166); -lean_ctor_set(x_170, 3, x_169); -lean_ctor_set(x_170, 4, x_168); -lean_ctor_set_uint8(x_170, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_170, sizeof(void*)*5 + 1, x_8); -x_171 = lean_st_ref_take(x_5, x_164); -x_172 = lean_ctor_get(x_171, 0); -lean_inc(x_172); -x_173 = lean_ctor_get(x_171, 1); -lean_inc(x_173); -lean_dec(x_171); -x_174 = lean_ctor_get(x_172, 0); -lean_inc(x_174); -x_175 = lean_ctor_get(x_172, 1); -lean_inc(x_175); -x_176 = lean_ctor_get(x_172, 2); -lean_inc(x_176); -x_177 = lean_ctor_get(x_172, 3); -lean_inc(x_177); -x_178 = lean_ctor_get(x_172, 4); -lean_inc(x_178); -x_179 = lean_ctor_get(x_172, 5); -lean_inc(x_179); -x_180 = lean_ctor_get(x_172, 6); -lean_inc(x_180); -if (lean_is_exclusive(x_172)) { - lean_ctor_release(x_172, 0); - lean_ctor_release(x_172, 1); - lean_ctor_release(x_172, 2); - lean_ctor_release(x_172, 3); - lean_ctor_release(x_172, 4); - lean_ctor_release(x_172, 5); - lean_ctor_release(x_172, 6); - x_181 = x_172; -} else { - lean_dec_ref(x_172); - x_181 = lean_box(0); -} -x_182 = l_Lean_PersistentArray_push___rarg(x_179, x_170); -if (lean_is_scalar(x_181)) { - x_183 = lean_alloc_ctor(0, 7, 0); -} else { - x_183 = x_181; -} -lean_ctor_set(x_183, 0, x_174); -lean_ctor_set(x_183, 1, x_175); -lean_ctor_set(x_183, 2, x_176); -lean_ctor_set(x_183, 3, x_177); -lean_ctor_set(x_183, 4, x_178); -lean_ctor_set(x_183, 5, x_182); -lean_ctor_set(x_183, 6, x_180); -x_184 = lean_st_ref_set(x_5, x_183, x_173); -x_185 = lean_ctor_get(x_184, 1); -lean_inc(x_185); -if (lean_is_exclusive(x_184)) { - lean_ctor_release(x_184, 0); - lean_ctor_release(x_184, 1); - x_186 = x_184; -} else { - lean_dec_ref(x_184); - x_186 = lean_box(0); -} -x_187 = lean_box(0); -if (lean_is_scalar(x_186)) { - x_188 = lean_alloc_ctor(0, 2, 0); -} else { - x_188 = x_186; +uint8_t x_16; +x_16 = 0; +return x_16; } -lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_185); -return x_188; } +else +{ +uint8_t x_17; +lean_inc(x_3); +x_17 = l_Lean_isRecCore(x_1, x_3); +if (x_17 == 0) +{ +uint8_t x_18; +lean_dec(x_3); +x_18 = 0; +return x_18; } else { -lean_object* x_189; uint8_t x_190; -x_189 = lean_ctor_get(x_16, 0); -lean_inc(x_189); -lean_dec(x_16); -x_190 = !lean_is_exclusive(x_17); -if (x_190 == 0) +lean_object* x_19; uint8_t x_20; +x_19 = l___private_Lean_CoreM_0__Lean_supportedRecursors; +x_20 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_19, x_3); +lean_dec(x_3); +if (x_20 == 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; 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; uint8_t x_204; -x_191 = lean_ctor_get(x_17, 0); -x_192 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); -lean_inc(x_194); -lean_dec(x_192); -x_195 = l_Lean_FileMap_toPosition(x_10, x_189); -lean_dec(x_189); -x_196 = l_Lean_FileMap_toPosition(x_10, x_191); -lean_dec(x_191); -lean_ctor_set(x_17, 0, x_196); -lean_inc(x_13); -lean_inc(x_12); -x_197 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_197, 0, x_12); -lean_ctor_set(x_197, 1, x_13); -x_198 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_198, 0, x_197); -lean_ctor_set(x_198, 1, x_193); -x_199 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_200 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_200, 0, x_9); -lean_ctor_set(x_200, 1, x_195); -lean_ctor_set(x_200, 2, x_17); -lean_ctor_set(x_200, 3, x_199); -lean_ctor_set(x_200, 4, x_198); -lean_ctor_set_uint8(x_200, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_200, sizeof(void*)*5 + 1, x_8); -x_201 = lean_st_ref_take(x_5, x_194); -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_is_exclusive(x_202); -if (x_204 == 0) -{ -lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; -x_205 = lean_ctor_get(x_202, 5); -x_206 = l_Lean_PersistentArray_push___rarg(x_205, x_200); -lean_ctor_set(x_202, 5, x_206); -x_207 = lean_st_ref_set(x_5, x_202, x_203); -x_208 = !lean_is_exclusive(x_207); -if (x_208 == 0) -{ -lean_object* x_209; lean_object* x_210; -x_209 = lean_ctor_get(x_207, 0); -lean_dec(x_209); -x_210 = lean_box(0); -lean_ctor_set(x_207, 0, x_210); -return x_207; -} -else -{ -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_207, 1); -lean_inc(x_211); -lean_dec(x_207); -x_212 = lean_box(0); -x_213 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_213, 0, x_212); -lean_ctor_set(x_213, 1, x_211); -return x_213; -} -} -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; -x_214 = lean_ctor_get(x_202, 0); -x_215 = lean_ctor_get(x_202, 1); -x_216 = lean_ctor_get(x_202, 2); -x_217 = lean_ctor_get(x_202, 3); -x_218 = lean_ctor_get(x_202, 4); -x_219 = lean_ctor_get(x_202, 5); -x_220 = lean_ctor_get(x_202, 6); -lean_inc(x_220); -lean_inc(x_219); -lean_inc(x_218); -lean_inc(x_217); -lean_inc(x_216); -lean_inc(x_215); -lean_inc(x_214); -lean_dec(x_202); -x_221 = l_Lean_PersistentArray_push___rarg(x_219, x_200); -x_222 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_222, 0, x_214); -lean_ctor_set(x_222, 1, x_215); -lean_ctor_set(x_222, 2, x_216); -lean_ctor_set(x_222, 3, x_217); -lean_ctor_set(x_222, 4, x_218); -lean_ctor_set(x_222, 5, x_221); -lean_ctor_set(x_222, 6, x_220); -x_223 = lean_st_ref_set(x_5, x_222, x_203); -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_223); - x_225 = lean_box(0); +uint8_t x_21; +x_21 = 1; +return x_21; +} +else +{ +uint8_t x_22; +x_22 = 0; +return x_22; +} } -x_226 = lean_box(0); -if (lean_is_scalar(x_225)) { - x_227 = lean_alloc_ctor(0, 2, 0); -} else { - x_227 = x_225; } -lean_ctor_set(x_227, 0, x_226); -lean_ctor_set(x_227, 1, x_224); -return x_227; } } else { -lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; -x_228 = lean_ctor_get(x_17, 0); -lean_inc(x_228); -lean_dec(x_17); -x_229 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); -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 = l_Lean_FileMap_toPosition(x_10, x_189); -lean_dec(x_189); -x_233 = l_Lean_FileMap_toPosition(x_10, x_228); -lean_dec(x_228); -x_234 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_234, 0, x_233); -lean_inc(x_13); -lean_inc(x_12); -x_235 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_235, 0, x_12); -lean_ctor_set(x_235, 1, x_13); -x_236 = lean_alloc_ctor(4, 2, 0); -lean_ctor_set(x_236, 0, x_235); -lean_ctor_set(x_236, 1, x_230); -x_237 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -lean_inc(x_9); -x_238 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_238, 0, x_9); -lean_ctor_set(x_238, 1, x_232); -lean_ctor_set(x_238, 2, x_234); -lean_ctor_set(x_238, 3, x_237); -lean_ctor_set(x_238, 4, x_236); -lean_ctor_set_uint8(x_238, sizeof(void*)*5, x_15); -lean_ctor_set_uint8(x_238, sizeof(void*)*5 + 1, x_8); -x_239 = lean_st_ref_take(x_5, x_231); -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 = lean_ctor_get(x_240, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_240, 1); -lean_inc(x_243); -x_244 = lean_ctor_get(x_240, 2); -lean_inc(x_244); -x_245 = lean_ctor_get(x_240, 3); -lean_inc(x_245); -x_246 = lean_ctor_get(x_240, 4); -lean_inc(x_246); -x_247 = lean_ctor_get(x_240, 5); -lean_inc(x_247); -x_248 = lean_ctor_get(x_240, 6); -lean_inc(x_248); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - lean_ctor_release(x_240, 2); - lean_ctor_release(x_240, 3); - lean_ctor_release(x_240, 4); - lean_ctor_release(x_240, 5); - lean_ctor_release(x_240, 6); - x_249 = x_240; -} else { - lean_dec_ref(x_240); - x_249 = lean_box(0); +uint8_t x_23; +lean_dec(x_2); +lean_dec(x_1); +x_23 = 0; +return x_23; } -x_250 = l_Lean_PersistentArray_push___rarg(x_247, x_238); -if (lean_is_scalar(x_249)) { - x_251 = lean_alloc_ctor(0, 7, 0); -} else { - x_251 = x_249; -} -lean_ctor_set(x_251, 0, x_242); -lean_ctor_set(x_251, 1, x_243); -lean_ctor_set(x_251, 2, x_244); -lean_ctor_set(x_251, 3, x_245); -lean_ctor_set(x_251, 4, x_246); -lean_ctor_set(x_251, 5, x_250); -lean_ctor_set(x_251, 6, x_248); -x_252 = lean_st_ref_set(x_5, x_251, x_241); -x_253 = lean_ctor_get(x_252, 1); -lean_inc(x_253); -if (lean_is_exclusive(x_252)) { - lean_ctor_release(x_252, 0); - lean_ctor_release(x_252, 1); - x_254 = x_252; -} else { - lean_dec_ref(x_252); - x_254 = lean_box(0); } -x_255 = lean_box(0); -if (lean_is_scalar(x_254)) { - x_256 = lean_alloc_ctor(0, 2, 0); -} else { - x_256 = x_254; } -lean_ctor_set(x_256, 0, x_255); -lean_ctor_set(x_256, 1, x_253); -return x_256; +static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("code generator does not support recursor '", 42); +return x_1; +} } +static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } +static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("' yet, consider using 'match ... with' and/or structural recursion", 66); +return x_1; } } +static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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) { _start: { lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_3, 5); +x_6 = lean_ctor_get(x_1, 1); lean_inc(x_6); -x_7 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_6, x_1, x_2, x_3, x_4, x_5); +lean_dec(x_1); +x_7 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x3f___spec__1(x_2, x_6); +if (lean_obj_tag(x_7) == 0) +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_dec(x_4); +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); lean_dec(x_3); -lean_dec(x_6); -return x_7; -} +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); +return x_11; } -LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_addDecl___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: +else { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 2); -lean_inc(x_5); -x_6 = l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1; -x_7 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_5, x_6); -lean_dec(x_5); -if (x_7 == 0) +lean_object* x_12; +x_12 = lean_ctor_get(x_7, 0); +lean_inc(x_12); +lean_dec(x_7); +if (lean_obj_tag(x_12) == 4) { -uint8_t x_8; lean_object* x_9; -x_8 = 1; -x_9 = l_Lean_log___at_Lean_addDecl___spec__5(x_1, x_8, x_2, x_3, x_4); -return x_9; +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_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_MessageData_ofName(x_13); +x_15 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2; +x_16 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_16, 0, x_15); +lean_ctor_set(x_16, 1, x_14); +x_17 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4; +x_18 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_18, 0, x_16); +lean_ctor_set(x_18, 1, x_17); +x_19 = l_Lean_throwError___rarg(x_3, x_4, x_18); +return x_19; } else { -uint8_t x_10; lean_object* x_11; -x_10 = 2; -x_11 = l_Lean_log___at_Lean_addDecl___spec__5(x_1, x_10, x_2, x_3, x_4); -return x_11; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +lean_dec(x_12); +lean_dec(x_4); +x_20 = lean_ctor_get(x_3, 0); +lean_inc(x_20); +lean_dec(x_3); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_box(0); +x_23 = lean_apply_2(x_21, lean_box(0), x_22); +return x_23; } } } -static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1() { +} +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___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_1; -x_1 = l_Lean_inheritedTraceOptions; -return x_1; +lean_object* x_6; +x_6 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(x_1, x_2, x_3, x_5, x_4); +return x_6; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_5; lean_object* x_6; uint8_t x_7; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1; -x_6 = lean_st_ref_get(x_5, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) +if (lean_obj_tag(x_5) == 0) { -lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; -x_8 = lean_ctor_get(x_6, 0); -x_9 = lean_ctor_get(x_2, 2); -x_10 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_8, x_9, x_1); -lean_dec(x_8); -x_11 = lean_box(x_10); -lean_ctor_set(x_6, 0, x_11); -return x_6; +lean_object* x_6; lean_object* x_7; lean_object* x_8; +lean_dec(x_3); +lean_dec(x_2); +x_6 = lean_ctor_get(x_1, 0); +lean_inc(x_6); +lean_dec(x_1); +x_7 = lean_ctor_get(x_6, 1); +lean_inc(x_7); +lean_dec(x_6); +x_8 = lean_apply_2(x_7, lean_box(0), x_4); +return x_8; } else { -lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; -x_12 = lean_ctor_get(x_6, 0); -x_13 = lean_ctor_get(x_6, 1); -lean_inc(x_13); +lean_object* x_9; lean_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_dec(x_4); +x_9 = lean_ctor_get(x_5, 0); +lean_inc(x_9); +x_10 = lean_ctor_get(x_5, 1); +lean_inc(x_10); +lean_dec(x_5); +x_11 = lean_ctor_get(x_1, 1); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 0); lean_inc(x_12); -lean_dec(x_6); -x_14 = lean_ctor_get(x_2, 2); -x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_14, x_1); -lean_dec(x_12); -x_16 = lean_box(x_15); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_16); -lean_ctor_set(x_17, 1, x_13); -return x_17; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; -x_3 = lean_st_ref_get(x_1, x_2); -x_4 = lean_ctor_get(x_3, 0); -lean_inc(x_4); -x_5 = lean_ctor_get(x_3, 1); -lean_inc(x_5); -lean_dec(x_3); -x_6 = lean_ctor_get(x_4, 3); -lean_inc(x_6); -lean_dec(x_4); -x_7 = lean_st_ref_take(x_1, x_5); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_ctor_get(x_7, 1); -lean_inc(x_9); -lean_dec(x_7); -x_10 = !lean_is_exclusive(x_8); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_11 = lean_ctor_get(x_8, 3); -lean_dec(x_11); -x_12 = l_Lean_Core_State_traceState___default___closed__3; -lean_ctor_set(x_8, 3, x_12); -x_13 = lean_st_ref_set(x_1, x_8, x_9); -x_14 = !lean_is_exclusive(x_13); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_13, 0); -lean_dec(x_15); -lean_ctor_set(x_13, 0, x_6); -return x_13; -} -else -{ -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_6); -lean_ctor_set(x_17, 1, x_16); -return x_17; -} -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_18 = lean_ctor_get(x_8, 0); -x_19 = lean_ctor_get(x_8, 1); -x_20 = lean_ctor_get(x_8, 2); -x_21 = lean_ctor_get(x_8, 4); -x_22 = lean_ctor_get(x_8, 5); -x_23 = lean_ctor_get(x_8, 6); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_8); -x_24 = l_Lean_Core_State_traceState___default___closed__3; -x_25 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_25, 0, x_18); -lean_ctor_set(x_25, 1, x_19); -lean_ctor_set(x_25, 2, x_20); -lean_ctor_set(x_25, 3, x_24); -lean_ctor_set(x_25, 4, x_21); -lean_ctor_set(x_25, 5, x_22); -lean_ctor_set(x_25, 6, x_23); -x_26 = lean_st_ref_set(x_1, x_25, x_9); -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -if (lean_is_exclusive(x_26)) { - lean_ctor_release(x_26, 0); - lean_ctor_release(x_26, 1); - x_28 = x_26; -} else { - lean_dec_ref(x_26); - x_28 = lean_box(0); -} -if (lean_is_scalar(x_28)) { - x_29 = lean_alloc_ctor(0, 2, 0); -} else { - x_29 = x_28; -} -lean_ctor_set(x_29, 0, x_6); -lean_ctor_set(x_29, 1, x_27); -return x_29; -} -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed), 2, 0); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___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) { -_start: -{ -uint8_t x_8; -x_8 = !lean_is_exclusive(x_5); -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; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; 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_9 = lean_ctor_get(x_5, 5); -x_10 = l_Lean_replaceRef(x_3, x_9); -lean_dec(x_9); -lean_ctor_set(x_5, 5, x_10); -x_11 = lean_st_ref_get(x_6, x_7); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 3); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_PersistentArray_toArray___rarg(x_14); -x_16 = lean_array_get_size(x_15); -x_17 = lean_usize_of_nat(x_16); -lean_dec(x_16); -x_18 = 0; -x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_17, x_18, x_15); -x_20 = lean_alloc_ctor(9, 3, 0); -lean_ctor_set(x_20, 0, x_2); -lean_ctor_set(x_20, 1, x_4); -lean_ctor_set(x_20, 2, x_19); -x_21 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_20, x_5, x_6, x_13); -lean_dec(x_5); -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_st_ref_take(x_6, x_23); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = !lean_is_exclusive(x_25); -if (x_27 == 0) -{ -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_25, 3); -lean_dec(x_28); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_3); -lean_ctor_set(x_29, 1, x_22); -x_30 = l_Lean_PersistentArray_push___rarg(x_1, x_29); -lean_ctor_set(x_25, 3, x_30); -x_31 = lean_st_ref_set(x_6, x_25, x_26); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -x_34 = lean_box(0); -lean_ctor_set(x_31, 0, x_34); -return x_31; -} -else -{ -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_35); -lean_dec(x_31); -x_36 = lean_box(0); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -return x_37; -} -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_38 = lean_ctor_get(x_25, 0); -x_39 = lean_ctor_get(x_25, 1); -x_40 = lean_ctor_get(x_25, 2); -x_41 = lean_ctor_get(x_25, 4); -x_42 = lean_ctor_get(x_25, 5); -x_43 = lean_ctor_get(x_25, 6); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_25); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_3); -lean_ctor_set(x_44, 1, x_22); -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_38); -lean_ctor_set(x_46, 1, x_39); -lean_ctor_set(x_46, 2, x_40); -lean_ctor_set(x_46, 3, x_45); -lean_ctor_set(x_46, 4, x_41); -lean_ctor_set(x_46, 5, x_42); -lean_ctor_set(x_46, 6, x_43); -x_47 = lean_st_ref_set(x_6, x_46, x_26); -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; -} -} -else -{ -lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; uint8_t 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_52 = lean_ctor_get(x_5, 0); -x_53 = lean_ctor_get(x_5, 1); -x_54 = lean_ctor_get(x_5, 2); -x_55 = lean_ctor_get(x_5, 3); -x_56 = lean_ctor_get(x_5, 4); -x_57 = lean_ctor_get(x_5, 5); -x_58 = lean_ctor_get(x_5, 6); -x_59 = lean_ctor_get(x_5, 7); -x_60 = lean_ctor_get(x_5, 8); -x_61 = lean_ctor_get(x_5, 9); -x_62 = lean_ctor_get(x_5, 10); -x_63 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); -x_64 = lean_ctor_get_uint8(x_5, sizeof(void*)*11 + 1); -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_inc(x_52); -lean_dec(x_5); -x_65 = l_Lean_replaceRef(x_3, x_57); -lean_dec(x_57); -x_66 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_66, 0, x_52); -lean_ctor_set(x_66, 1, x_53); -lean_ctor_set(x_66, 2, x_54); -lean_ctor_set(x_66, 3, x_55); -lean_ctor_set(x_66, 4, x_56); -lean_ctor_set(x_66, 5, x_65); -lean_ctor_set(x_66, 6, x_58); -lean_ctor_set(x_66, 7, x_59); -lean_ctor_set(x_66, 8, x_60); -lean_ctor_set(x_66, 9, x_61); -lean_ctor_set(x_66, 10, x_62); -lean_ctor_set_uint8(x_66, sizeof(void*)*11, x_63); -lean_ctor_set_uint8(x_66, sizeof(void*)*11 + 1, x_64); -x_67 = lean_st_ref_get(x_6, x_7); -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, 0); -lean_ctor_set(x_76, 0, x_2); -lean_ctor_set(x_76, 1, x_4); -lean_ctor_set(x_76, 2, x_75); -x_77 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_76, x_66, x_6, 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_6, 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_6, 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__11(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_5; lean_object* x_6; -x_5 = lean_ctor_get(x_1, 0); -lean_inc(x_5); -x_6 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_4); -return x_6; -} -else -{ -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_1, 0); -lean_inc(x_7); -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_7); -lean_ctor_set(x_8, 1, x_4); -return x_8; -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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_inc(x_7); -x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10(x_1, x_5, x_2, x_3, x_7, x_8, x_9); -x_11 = lean_ctor_get(x_10, 1); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(x_4, x_7, x_8, x_11); -lean_dec(x_7); -lean_dec(x_4); -return x_12; -} -} -static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; uint8_t x_2; double x_3; -x_1 = lean_unsigned_to_nat(0u); -x_2 = 0; -x_3 = l_Float_ofScientific(x_1, x_2, x_1); -return x_3; -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_profiler; -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -double x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1; -lean_inc(x_3); -lean_inc(x_1); -x_16 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_16, 0, x_1); -lean_ctor_set(x_16, 1, x_3); -lean_ctor_set_float(x_16, sizeof(void*)*2, x_15); -lean_ctor_set_float(x_16, sizeof(void*)*2 + 8, x_15); -lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 16, x_2); -x_17 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2; -x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_17); -lean_dec(x_7); -if (x_18 == 0) -{ -if (x_8 == 0) -{ -lean_object* x_19; lean_object* x_20; -lean_dec(x_3); -lean_dec(x_1); -x_19 = lean_box(0); -x_20 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_16, x_19, x_12, x_13, x_14); -return x_20; -} -else -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_16); -x_21 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_21, 0, x_1); -lean_ctor_set(x_21, 1, x_3); -lean_ctor_set_float(x_21, sizeof(void*)*2, x_9); -lean_ctor_set_float(x_21, sizeof(void*)*2 + 8, x_10); -lean_ctor_set_uint8(x_21, sizeof(void*)*2 + 16, x_2); -x_22 = lean_box(0); -x_23 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_21, x_22, x_12, x_13, x_14); -return x_23; -} -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; -lean_dec(x_16); -x_24 = lean_alloc_ctor(0, 2, 17); -lean_ctor_set(x_24, 0, x_1); -lean_ctor_set(x_24, 1, x_3); -lean_ctor_set_float(x_24, sizeof(void*)*2, x_9); -lean_ctor_set_float(x_24, sizeof(void*)*2 + 8, x_10); -lean_ctor_set_uint8(x_24, sizeof(void*)*2 + 16, x_2); -x_25 = lean_box(0); -x_26 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_4, x_5, x_11, x_6, x_24, x_25, x_12, x_13, x_14); -return x_26; -} -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 53); -return x_1; -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, uint8_t x_7, double x_8, double x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; lean_object* x_16; -lean_dec(x_11); -x_15 = lean_ctor_get(x_12, 5); -lean_inc(x_15); -lean_inc(x_13); -lean_inc(x_12); -lean_inc(x_5); -x_16 = lean_apply_4(x_10, x_5, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_1, x_2, x_3, x_4, x_15, x_5, x_6, x_7, x_8, x_9, x_17, x_12, x_13, x_18); -lean_dec(x_13); -return x_19; -} -else -{ -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 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; -x_22 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_1, x_2, x_3, x_4, x_15, x_5, x_6, x_7, x_8, x_9, x_21, x_12, x_13, x_20); -lean_dec(x_13); -return x_22; -} -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_trace_profiler_useHearbeats; -return x_1; -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_trace_profiler; -return x_1; -} -} -static lean_object* _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = l_Lean_trace_profiler_threshold; -return x_1; -} -} -static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -static double _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5() { -_start: -{ -lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; -x_1 = lean_unsigned_to_nat(1000000000u); -x_2 = 0; -x_3 = lean_unsigned_to_nat(0u); -x_4 = l_Float_ofScientific(x_1, x_2, x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___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; uint8_t x_16; -lean_dec(x_8); -x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___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); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; -x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_85 = lean_io_mono_nanos_now(x_14); -x_86 = lean_ctor_get(x_85, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_85, 1); -lean_inc(x_87); -lean_dec(x_85); -lean_inc(x_10); -lean_inc(x_9); -x_88 = lean_apply_3(x_7, x_9, x_10, x_87); -if (lean_obj_tag(x_88) == 0) -{ -lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; uint8_t x_95; lean_object* x_96; double x_97; double x_98; double x_99; double x_100; double x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -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_alloc_ctor(1, 1, 0); -lean_ctor_set(x_91, 0, x_89); -x_92 = lean_io_mono_nanos_now(x_90); -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 = 0; -x_96 = lean_unsigned_to_nat(0u); -x_97 = l_Float_ofScientific(x_86, x_95, x_96); -lean_dec(x_86); -x_98 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5; -x_99 = lean_float_div(x_97, x_98); -x_100 = l_Float_ofScientific(x_93, x_95, x_96); -lean_dec(x_93); -x_101 = lean_float_div(x_100, x_98); -x_102 = lean_box_float(x_99); -x_103 = lean_box_float(x_101); -x_104 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_104, 0, x_102); -lean_ctor_set(x_104, 1, x_103); -x_105 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_105, 0, x_91); -lean_ctor_set(x_105, 1, x_104); -x_17 = x_105; -x_18 = x_94; -goto block_84; -} -else -{ -lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; double x_114; double x_115; double x_116; double x_117; double x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; -x_106 = lean_ctor_get(x_88, 0); -lean_inc(x_106); -x_107 = lean_ctor_get(x_88, 1); -lean_inc(x_107); -lean_dec(x_88); -x_108 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_108, 0, x_106); -x_109 = lean_io_mono_nanos_now(x_107); -x_110 = lean_ctor_get(x_109, 0); -lean_inc(x_110); -x_111 = lean_ctor_get(x_109, 1); -lean_inc(x_111); -lean_dec(x_109); -x_112 = 0; -x_113 = lean_unsigned_to_nat(0u); -x_114 = l_Float_ofScientific(x_86, x_112, x_113); -lean_dec(x_86); -x_115 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5; -x_116 = lean_float_div(x_114, x_115); -x_117 = l_Float_ofScientific(x_110, x_112, x_113); -lean_dec(x_110); -x_118 = lean_float_div(x_117, x_115); -x_119 = lean_box_float(x_116); -x_120 = lean_box_float(x_118); -x_121 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -x_122 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_122, 0, x_108); -lean_ctor_set(x_122, 1, x_121); -x_17 = x_122; -x_18 = x_111; -goto block_84; -} -block_84: -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_70; uint8_t x_71; -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -x_20 = lean_ctor_get(x_17, 0); -lean_inc(x_20); -lean_dec(x_17); -x_21 = lean_ctor_get(x_19, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_19, 1); -lean_inc(x_22); -lean_dec(x_19); -x_70 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; -x_71 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_70); -if (x_71 == 0) -{ -uint8_t x_72; -x_72 = 0; -x_23 = x_72; -goto block_69; -} -else -{ -double x_73; double x_74; double x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; double x_80; double x_81; double x_82; uint8_t x_83; -x_73 = lean_unbox_float(x_22); -x_74 = lean_unbox_float(x_21); -x_75 = lean_float_sub(x_73, x_74); -x_76 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3; -x_77 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_76); -x_78 = 0; -x_79 = lean_unsigned_to_nat(0u); -x_80 = l_Float_ofScientific(x_77, x_78, x_79); -lean_dec(x_77); -x_81 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4; -x_82 = lean_float_div(x_80, x_81); -x_83 = lean_float_decLt(x_82, x_75); -x_23 = x_83; -goto block_69; -} -block_69: -{ -if (x_6 == 0) -{ -if (x_23 == 0) -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; -lean_dec(x_22); -lean_dec(x_21); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_24 = lean_st_ref_take(x_10, x_18); -x_25 = lean_ctor_get(x_24, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -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, 3); -x_29 = l_Lean_PersistentArray_append___rarg(x_13, x_28); -lean_ctor_set(x_25, 3, x_29); -x_30 = lean_st_ref_set(x_10, 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_addDecl___spec__11(x_20, x_9, x_10, x_31); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_20); -if (lean_obj_tag(x_32) == 0) -{ -uint8_t x_33; -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; -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) -{ -return x_32; -} -else -{ -lean_object* x_38; lean_object* x_39; lean_object* x_40; -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_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; -} -} -} -else -{ -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; 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_41 = lean_ctor_get(x_25, 0); -x_42 = lean_ctor_get(x_25, 1); -x_43 = lean_ctor_get(x_25, 2); -x_44 = lean_ctor_get(x_25, 3); -x_45 = lean_ctor_get(x_25, 4); -x_46 = lean_ctor_get(x_25, 5); -x_47 = lean_ctor_get(x_25, 6); -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_inc(x_41); -lean_dec(x_25); -x_48 = l_Lean_PersistentArray_append___rarg(x_13, x_44); -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_45); -lean_ctor_set(x_49, 5, x_46); -lean_ctor_set(x_49, 6, x_47); -x_50 = lean_st_ref_set(x_10, 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_addDecl___spec__11(x_20, x_9, x_10, x_51); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_20); -if (lean_obj_tag(x_52) == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_52, 0); -lean_inc(x_53); -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_52); - x_55 = lean_box(0); -} -if (lean_is_scalar(x_55)) { - x_56 = lean_alloc_ctor(0, 2, 0); -} else { - x_56 = x_55; -} -lean_ctor_set(x_56, 0, x_53); -lean_ctor_set(x_56, 1, x_54); -return x_56; -} -else -{ -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_57 = lean_ctor_get(x_52, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_52, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_52)) { - lean_ctor_release(x_52, 0); - lean_ctor_release(x_52, 1); - x_59 = x_52; -} else { - lean_dec_ref(x_52); - x_59 = lean_box(0); -} -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(1, 2, 0); -} else { - x_60 = x_59; -} -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -return x_60; -} -} -} -else -{ -lean_object* x_61; double x_62; double x_63; lean_object* x_64; -x_61 = lean_box(0); -x_62 = lean_unbox_float(x_21); -lean_dec(x_21); -x_63 = lean_unbox_float(x_22); -lean_dec(x_22); -x_64 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_3, x_4, x_13, x_20, x_1, x_23, x_62, x_63, x_5, x_61, x_9, x_10, x_18); -return x_64; -} -} -else -{ -lean_object* x_65; double x_66; double x_67; lean_object* x_68; -x_65 = lean_box(0); -x_66 = lean_unbox_float(x_21); -lean_dec(x_21); -x_67 = lean_unbox_float(x_22); -lean_dec(x_22); -x_68 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_3, x_4, x_13, x_20, x_1, x_23, x_66, x_67, x_5, x_65, x_9, x_10, x_18); -return x_68; -} -} -} -} -else -{ -lean_object* x_123; lean_object* x_124; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -x_189 = lean_io_get_num_heartbeats(x_14); -x_190 = lean_ctor_get(x_189, 0); -lean_inc(x_190); -x_191 = lean_ctor_get(x_189, 1); -lean_inc(x_191); -lean_dec(x_189); -lean_inc(x_10); -lean_inc(x_9); -x_192 = lean_apply_3(x_7, x_9, x_10, x_191); -if (lean_obj_tag(x_192) == 0) -{ -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; uint8_t x_199; lean_object* x_200; double x_201; double x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); -lean_inc(x_194); -lean_dec(x_192); -x_195 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_195, 0, x_193); -x_196 = lean_io_get_num_heartbeats(x_194); -x_197 = lean_ctor_get(x_196, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_196, 1); -lean_inc(x_198); -lean_dec(x_196); -x_199 = 0; -x_200 = lean_unsigned_to_nat(0u); -x_201 = l_Float_ofScientific(x_190, x_199, x_200); -lean_dec(x_190); -x_202 = l_Float_ofScientific(x_197, x_199, x_200); -lean_dec(x_197); -x_203 = lean_box_float(x_201); -x_204 = lean_box_float(x_202); -x_205 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_205, 0, x_203); -lean_ctor_set(x_205, 1, x_204); -x_206 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_206, 0, x_195); -lean_ctor_set(x_206, 1, x_205); -x_123 = x_206; -x_124 = x_198; -goto block_188; -} -else -{ -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_object* x_214; double x_215; double x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; -x_207 = lean_ctor_get(x_192, 0); -lean_inc(x_207); -x_208 = lean_ctor_get(x_192, 1); -lean_inc(x_208); -lean_dec(x_192); -x_209 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_209, 0, x_207); -x_210 = lean_io_get_num_heartbeats(x_208); -x_211 = lean_ctor_get(x_210, 0); -lean_inc(x_211); -x_212 = lean_ctor_get(x_210, 1); -lean_inc(x_212); -lean_dec(x_210); -x_213 = 0; -x_214 = lean_unsigned_to_nat(0u); -x_215 = l_Float_ofScientific(x_190, x_213, x_214); -lean_dec(x_190); -x_216 = l_Float_ofScientific(x_211, x_213, x_214); -lean_dec(x_211); -x_217 = lean_box_float(x_215); -x_218 = lean_box_float(x_216); -x_219 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_219, 0, x_217); -lean_ctor_set(x_219, 1, x_218); -x_220 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_220, 0, x_209); -lean_ctor_set(x_220, 1, x_219); -x_123 = x_220; -x_124 = x_212; -goto block_188; -} -block_188: -{ -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_176; uint8_t x_177; -x_125 = lean_ctor_get(x_123, 1); -lean_inc(x_125); -x_126 = lean_ctor_get(x_123, 0); -lean_inc(x_126); -lean_dec(x_123); -x_127 = lean_ctor_get(x_125, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_125, 1); -lean_inc(x_128); -lean_dec(x_125); -x_176 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; -x_177 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_176); -if (x_177 == 0) -{ -uint8_t x_178; -x_178 = 0; -x_129 = x_178; -goto block_175; -} -else -{ -double x_179; double x_180; double x_181; lean_object* x_182; lean_object* x_183; uint8_t x_184; lean_object* x_185; double x_186; uint8_t x_187; -x_179 = lean_unbox_float(x_128); -x_180 = lean_unbox_float(x_127); -x_181 = lean_float_sub(x_179, x_180); -x_182 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3; -x_183 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_182); -x_184 = 0; -x_185 = lean_unsigned_to_nat(0u); -x_186 = l_Float_ofScientific(x_183, x_184, x_185); -lean_dec(x_183); -x_187 = lean_float_decLt(x_186, x_181); -x_129 = x_187; -goto block_175; -} -block_175: -{ -if (x_6 == 0) -{ -if (x_129 == 0) -{ -lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; -lean_dec(x_128); -lean_dec(x_127); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_2); -lean_dec(x_1); -x_130 = lean_st_ref_take(x_10, x_124); -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_130, 1); -lean_inc(x_132); -lean_dec(x_130); -x_133 = !lean_is_exclusive(x_131); -if (x_133 == 0) -{ -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; -x_134 = lean_ctor_get(x_131, 3); -x_135 = l_Lean_PersistentArray_append___rarg(x_13, x_134); -lean_ctor_set(x_131, 3, x_135); -x_136 = lean_st_ref_set(x_10, x_131, x_132); -x_137 = lean_ctor_get(x_136, 1); -lean_inc(x_137); -lean_dec(x_136); -x_138 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(x_126, x_9, x_10, x_137); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_126); -if (lean_obj_tag(x_138) == 0) -{ -uint8_t x_139; -x_139 = !lean_is_exclusive(x_138); -if (x_139 == 0) -{ -return x_138; -} -else -{ -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_138, 0); -x_141 = lean_ctor_get(x_138, 1); -lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_138); -x_142 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -return x_142; -} -} -else -{ -uint8_t x_143; -x_143 = !lean_is_exclusive(x_138); -if (x_143 == 0) -{ -return x_138; -} -else -{ -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_138, 0); -x_145 = lean_ctor_get(x_138, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_138); -x_146 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; -} -} -} -else -{ -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_147 = lean_ctor_get(x_131, 0); -x_148 = lean_ctor_get(x_131, 1); -x_149 = lean_ctor_get(x_131, 2); -x_150 = lean_ctor_get(x_131, 3); -x_151 = lean_ctor_get(x_131, 4); -x_152 = lean_ctor_get(x_131, 5); -x_153 = lean_ctor_get(x_131, 6); -lean_inc(x_153); -lean_inc(x_152); -lean_inc(x_151); -lean_inc(x_150); -lean_inc(x_149); -lean_inc(x_148); -lean_inc(x_147); -lean_dec(x_131); -x_154 = l_Lean_PersistentArray_append___rarg(x_13, x_150); -x_155 = lean_alloc_ctor(0, 7, 0); -lean_ctor_set(x_155, 0, x_147); -lean_ctor_set(x_155, 1, x_148); -lean_ctor_set(x_155, 2, x_149); -lean_ctor_set(x_155, 3, x_154); -lean_ctor_set(x_155, 4, x_151); -lean_ctor_set(x_155, 5, x_152); -lean_ctor_set(x_155, 6, x_153); -x_156 = lean_st_ref_set(x_10, x_155, x_132); -x_157 = lean_ctor_get(x_156, 1); -lean_inc(x_157); -lean_dec(x_156); -x_158 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(x_126, x_9, x_10, x_157); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_126); -if (lean_obj_tag(x_158) == 0) -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_159 = lean_ctor_get(x_158, 0); -lean_inc(x_159); -x_160 = lean_ctor_get(x_158, 1); -lean_inc(x_160); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_161 = x_158; -} else { - lean_dec_ref(x_158); - x_161 = lean_box(0); -} -if (lean_is_scalar(x_161)) { - x_162 = lean_alloc_ctor(0, 2, 0); -} else { - x_162 = x_161; -} -lean_ctor_set(x_162, 0, x_159); -lean_ctor_set(x_162, 1, x_160); -return x_162; -} -else -{ -lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_163 = lean_ctor_get(x_158, 0); -lean_inc(x_163); -x_164 = lean_ctor_get(x_158, 1); -lean_inc(x_164); -if (lean_is_exclusive(x_158)) { - lean_ctor_release(x_158, 0); - lean_ctor_release(x_158, 1); - x_165 = x_158; -} else { - lean_dec_ref(x_158); - x_165 = lean_box(0); -} -if (lean_is_scalar(x_165)) { - x_166 = lean_alloc_ctor(1, 2, 0); -} else { - x_166 = x_165; -} -lean_ctor_set(x_166, 0, x_163); -lean_ctor_set(x_166, 1, x_164); -return x_166; -} -} -} -else -{ -lean_object* x_167; double x_168; double x_169; lean_object* x_170; -x_167 = lean_box(0); -x_168 = lean_unbox_float(x_127); -lean_dec(x_127); -x_169 = lean_unbox_float(x_128); -lean_dec(x_128); -x_170 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_3, x_4, x_13, x_126, x_1, x_129, x_168, x_169, x_5, x_167, x_9, x_10, x_124); -return x_170; -} -} -else -{ -lean_object* x_171; double x_172; double x_173; lean_object* x_174; -x_171 = lean_box(0); -x_172 = lean_unbox_float(x_127); -lean_dec(x_127); -x_173 = lean_unbox_float(x_128); -lean_dec(x_128); -x_174 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_2, x_3, x_4, x_13, x_126, x_1, x_129, x_172, x_173, x_5, x_171, x_9, x_10, x_124); -return x_174; -} -} -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_9 = lean_ctor_get(x_6, 2); -lean_inc(x_9); -lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_1, x_6, x_7, x_8); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_unbox(x_11); -if (x_12 == 0) -{ -lean_object* x_13; lean_object* x_14; uint8_t x_15; -x_13 = lean_ctor_get(x_10, 1); -lean_inc(x_13); -lean_dec(x_10); -x_14 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; -x_15 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_14); -if (x_15 == 0) -{ -lean_object* x_16; -lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_5); -lean_dec(x_2); -lean_dec(x_1); -x_16 = lean_apply_3(x_3, x_6, x_7, x_13); -if (lean_obj_tag(x_16) == 0) -{ -uint8_t x_17; -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) -{ -return x_16; -} -else -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_inc(x_18); -lean_dec(x_16); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_18); -lean_ctor_set(x_20, 1, x_19); -return x_20; -} -} -else -{ -uint8_t x_21; -x_21 = !lean_is_exclusive(x_16); -if (x_21 == 0) -{ -return x_16; -} -else -{ -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_16, 0); -x_23 = lean_ctor_get(x_16, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_16); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; -} -} -} -else -{ -lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_25 = lean_box(0); -x_26 = lean_unbox(x_11); -lean_dec(x_11); -x_27 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(x_9, x_1, x_4, x_5, x_2, x_26, x_3, x_25, x_6, x_7, x_13); -return x_27; -} -} -else -{ -lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_10, 1); -lean_inc(x_28); -lean_dec(x_10); -x_29 = lean_box(0); -x_30 = lean_unbox(x_11); -lean_dec(x_11); -x_31 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4(x_9, x_1, x_4, x_5, x_2, x_30, x_3, x_29, x_6, x_7, x_28); -return x_31; -} -} -} -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_apply_2(x_3, x_5, x_6); -x_9 = l_Lean_profileitIOUnsafe___rarg(x_1, x_2, x_8, x_4, x_7); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg___boxed), 7, 0); -return x_2; -} -} -static lean_object* _init_l_Lean_addDecl___lambda__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("typechecking declaration", 24); -return x_1; -} -} -static lean_object* _init_l_Lean_addDecl___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addDecl___lambda__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; -x_5 = l_Lean_addDecl___lambda__1___closed__2; -x_6 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_6, 0, x_5); -lean_ctor_set(x_6, 1, x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_dec(x_2); -x_6 = lean_st_ref_get(x_4, x_5); -x_7 = lean_ctor_get(x_6, 0); -lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = lean_ctor_get(x_7, 0); -lean_inc(x_9); -lean_dec(x_7); -x_10 = lean_add_decl(x_9, x_1); -lean_dec(x_1); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; lean_object* x_12; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -lean_dec(x_10); -x_12 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_11, x_3, x_4, x_8); -lean_dec(x_4); -return x_12; -} -else -{ -lean_object* x_13; lean_object* x_14; -x_13 = lean_ctor_get(x_10, 0); -lean_inc(x_13); -lean_dec(x_10); -x_14 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_13, x_3, x_4, x_8); -lean_dec(x_4); -lean_dec(x_3); -return x_14; -} -} -} -static lean_object* _init_l_Lean_addDecl___lambda__3___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("declaration uses 'sorry'", 24); -return x_1; -} -} -static lean_object* _init_l_Lean_addDecl___lambda__3___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addDecl___lambda__3___closed__1; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -static lean_object* _init_l_Lean_addDecl___lambda__3___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_addDecl___lambda__3___closed__2; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_5 = lean_st_ref_get(x_3, x_4); -x_6 = lean_ctor_get(x_5, 0); -lean_inc(x_6); -x_7 = lean_ctor_get(x_5, 1); -lean_inc(x_7); -lean_dec(x_5); -x_8 = lean_ctor_get(x_6, 5); -lean_inc(x_8); -lean_dec(x_6); -x_9 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_8); -if (x_9 == 0) -{ -uint8_t x_10; uint8_t x_11; -x_10 = 0; -lean_inc(x_1); -x_11 = l_Lean_Declaration_foldExprM___at_Lean_Declaration_hasSorry___spec__1(x_1, x_10); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; -x_12 = lean_box(0); -x_13 = l_Lean_addDecl___lambda__2(x_1, x_12, x_2, x_3, x_7); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_14 = l_Lean_addDecl___lambda__3___closed__3; -lean_inc(x_3); -lean_inc(x_2); -x_15 = l_Lean_logWarning___at_Lean_addDecl___spec__4(x_14, x_2, x_3, x_7); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); -x_18 = l_Lean_addDecl___lambda__2(x_1, x_16, x_2, x_3, x_17); -return x_18; -} -} -else -{ -lean_object* x_19; lean_object* x_20; -x_19 = lean_box(0); -x_20 = l_Lean_addDecl___lambda__2(x_1, x_19, x_2, x_3, x_7); -return x_20; -} -} -} -static lean_object* _init_l_Lean_addDecl___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__1___boxed), 4, 0); -return x_1; -} -} -static lean_object* _init_l_Lean_addDecl___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("type checking", 13); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_addDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; 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; -x_5 = lean_ctor_get(x_2, 2); -lean_inc(x_5); -x_6 = lean_alloc_closure((void*)(l_Lean_addDecl___lambda__3), 4, 1); -lean_closure_set(x_6, 0, x_1); -x_7 = l_Lean_Core_initFn____x40_Lean_CoreM___hyg_113____closed__2; -x_8 = l_Lean_addDecl___closed__1; -x_9 = 1; -x_10 = l_Lean_initFn____x40_Lean_CoreM___hyg_80____closed__3; -x_11 = lean_box(x_9); -x_12 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___boxed), 8, 5); -lean_closure_set(x_12, 0, x_7); -lean_closure_set(x_12, 1, x_8); -lean_closure_set(x_12, 2, x_6); -lean_closure_set(x_12, 3, x_11); -lean_closure_set(x_12, 4, x_10); -x_13 = l_Lean_addDecl___closed__2; -x_14 = lean_box(0); -x_15 = l_Lean_profileitM___at_Lean_addDecl___spec__12___rarg(x_13, x_5, x_12, x_14, x_2, x_3, x_4); -lean_dec(x_5); -return x_15; -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_addDecl___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -uint8_t x_7; lean_object* x_8; -x_7 = lean_unbox(x_3); -lean_dec(x_3); -x_8 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_1, x_2, x_7, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_log___at_Lean_addDecl___spec__5___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_6; lean_object* x_7; -x_6 = lean_unbox(x_2); -lean_dec(x_2); -x_7 = l_Lean_log___at_Lean_addDecl___spec__5(x_1, x_6, x_3, x_4, x_5); -return x_7; -} -} -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg___boxed(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(x_1, x_2); -lean_dec(x_1); -return x_3; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___boxed(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9(x_1); -lean_dec(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___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) { -_start: -{ -lean_object* x_8; -x_8 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -return x_8; -} -} -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_MonadExcept_ofExcept___at_Lean_addDecl___spec__11(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__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_withTraceNode___at_Lean_addDecl___spec__7___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_6); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; uint8_t x_16; double x_17; double x_18; lean_object* x_19; -x_15 = lean_unbox(x_2); -lean_dec(x_2); -x_16 = lean_unbox(x_8); -lean_dec(x_8); -x_17 = lean_unbox_float(x_9); -lean_dec(x_9); -x_18 = lean_unbox_float(x_10); -lean_dec(x_10); -x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_16, x_17, x_18, x_11, x_12, x_13, x_14); -lean_dec(x_13); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -uint8_t x_15; uint8_t x_16; double x_17; double x_18; lean_object* x_19; -x_15 = lean_unbox(x_2); -lean_dec(x_2); -x_16 = lean_unbox(x_7); -lean_dec(x_7); -x_17 = lean_unbox_float(x_8); -lean_dec(x_8); -x_18 = lean_unbox_float(x_9); -lean_dec(x_9); -x_19 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3(x_1, x_15, x_3, x_4, x_5, x_6, x_16, x_17, x_18, x_10, x_11, x_12, x_13, x_14); -return x_19; -} -} -LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { -_start: -{ -uint8_t x_12; uint8_t x_13; lean_object* x_14; -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_addDecl___spec__7___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_addDecl___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -uint8_t x_9; lean_object* x_10; -x_9 = lean_unbox(x_4); -lean_dec(x_4); -x_10 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7(x_1, x_2, x_3, x_9, x_5, x_6, x_7, x_8); -return x_10; -} -} -LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_addDecl___spec__12___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__12___rarg(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_2); -lean_dec(x_1); -return x_8; -} -} -LEAN_EXPORT lean_object* l_Lean_addDecl___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_addDecl___lambda__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_5; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Empty", 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("rec", 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("False", 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__5() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Eq", 2); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__7() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("ndrec", 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__7; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__9() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__10() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("recOn", 5); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__11() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__10; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("casesOn", 7); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__13() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__6; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__4; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("And", 3); -return x_1; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__17() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__16; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__12; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__19() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(10u); -x_2 = lean_mk_empty_array_with_capacity(x_1); -return x_2; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__19; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__3; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__20; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__5; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__22() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__21; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__8; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__22; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__9; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__24() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__23; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__11; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__25() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__24; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__13; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__25; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__14; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__26; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__15; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__27; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__17; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__28; -x_2 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__18; -x_3 = lean_array_push(x_1, x_2); -return x_3; -} -} -static lean_object* _init_l___private_Lean_CoreM_0__Lean_supportedRecursors() { -_start: -{ -lean_object* x_1; -x_1 = l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__29; -return x_1; -} -} -LEAN_EXPORT uint8_t l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_2) == 4) -{ -lean_object* x_3; uint8_t x_4; -x_3 = lean_ctor_get(x_2, 0); -lean_inc(x_3); -lean_dec(x_2); -lean_inc(x_3); -lean_inc(x_1); -x_4 = lean_is_aux_recursor(x_1, x_3); -if (x_4 == 0) -{ -uint8_t x_5; -lean_inc(x_3); -x_5 = l_Lean_isRecCore(x_1, x_3); -if (x_5 == 0) -{ -uint8_t x_6; -lean_dec(x_3); -x_6 = 0; -return x_6; -} -else -{ -lean_object* x_7; uint8_t x_8; -x_7 = l___private_Lean_CoreM_0__Lean_supportedRecursors; -x_8 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_7, x_3); -lean_dec(x_3); -if (x_8 == 0) -{ -uint8_t x_9; -x_9 = 1; -return x_9; -} -else -{ -uint8_t x_10; -x_10 = 0; -return x_10; -} -} -} -else -{ -lean_object* x_11; uint8_t x_12; -x_11 = l_Lean_casesOnSuffix; -lean_inc(x_3); -lean_inc(x_1); -x_12 = l_Lean_isAuxRecursorWithSuffix(x_1, x_3, x_11); -if (x_12 == 0) -{ -lean_object* x_13; uint8_t x_14; -lean_dec(x_1); -x_13 = l___private_Lean_CoreM_0__Lean_supportedRecursors; -x_14 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_13, x_3); -lean_dec(x_3); -if (x_14 == 0) -{ -uint8_t x_15; -x_15 = 1; -return x_15; -} -else -{ -uint8_t x_16; -x_16 = 0; -return x_16; -} -} -else -{ -uint8_t x_17; -lean_inc(x_3); -x_17 = l_Lean_isRecCore(x_1, x_3); -if (x_17 == 0) -{ -uint8_t x_18; -lean_dec(x_3); -x_18 = 0; -return x_18; -} -else -{ -lean_object* x_19; uint8_t x_20; -x_19 = l___private_Lean_CoreM_0__Lean_supportedRecursors; -x_20 = l_Array_contains___at_Lean_registerInternalExceptionId___spec__1(x_19, x_3); -lean_dec(x_3); -if (x_20 == 0) -{ -uint8_t x_21; -x_21 = 1; -return x_21; -} -else -{ -uint8_t x_22; -x_22 = 0; -return x_22; -} -} -} -} -} -else -{ -uint8_t x_23; -lean_dec(x_2); -lean_dec(x_1); -x_23 = 0; -return x_23; -} -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("code generator does not support recursor '", 42); -return x_1; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("' yet, consider using 'match ... with' and/or structural recursion", 66); -return x_1; -} -} -static lean_object* _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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) { -_start: -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -lean_dec(x_1); -x_7 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x3f___spec__1(x_2, x_6); -if (lean_obj_tag(x_7) == 0) -{ -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; -lean_dec(x_4); -x_8 = lean_ctor_get(x_3, 0); -lean_inc(x_8); -lean_dec(x_3); -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); -return x_11; -} -else -{ -lean_object* x_12; -x_12 = lean_ctor_get(x_7, 0); -lean_inc(x_12); -lean_dec(x_7); -if (lean_obj_tag(x_12) == 4) -{ -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_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -lean_dec(x_12); -x_14 = l_Lean_MessageData_ofName(x_13); -x_15 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__2; -x_16 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_16, 0, x_15); -lean_ctor_set(x_16, 1, x_14); -x_17 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4; -x_18 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_18, 1, x_17); -x_19 = l_Lean_throwError___rarg(x_3, x_4, x_18); -return x_19; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -lean_dec(x_12); -lean_dec(x_4); -x_20 = lean_ctor_get(x_3, 0); -lean_inc(x_20); -lean_dec(x_3); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_box(0); -x_23 = lean_apply_2(x_21, lean_box(0), x_22); -return x_23; -} -} -} -} -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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) { -_start: -{ -lean_object* x_6; -x_6 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(x_1, x_2, x_3, x_5, x_4); -return x_6; -} -} -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { -_start: -{ -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; lean_object* x_8; -lean_dec(x_3); -lean_dec(x_2); -x_6 = lean_ctor_get(x_1, 0); -lean_inc(x_6); -lean_dec(x_1); -x_7 = lean_ctor_get(x_6, 1); -lean_inc(x_7); -lean_dec(x_6); -x_8 = lean_apply_2(x_7, lean_box(0), x_4); -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_dec(x_4); -x_9 = lean_ctor_get(x_5, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_dec(x_5); -x_11 = lean_ctor_get(x_1, 1); -lean_inc(x_11); -x_12 = lean_ctor_get(x_9, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_12, 2); -lean_inc(x_13); +x_13 = lean_ctor_get(x_12, 2); +lean_inc(x_13); lean_dec(x_12); lean_inc(x_3); x_14 = lean_alloc_closure((void*)(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1___boxed), 2, 1); @@ -15257,169 +12375,587 @@ x_8 = lean_apply_4(x_5, lean_box(0), lean_box(0), x_6, x_7); return x_8; } } -LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported(lean_object* x_1) { -_start: -{ -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_CoreM_0__Lean_checkUnsupported___rarg), 4, 0); -return x_2; +LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_CoreM_0__Lean_checkUnsupported___rarg), 4, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1(x_1, x_2); +x_4 = lean_box(x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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) { +_start: +{ +lean_object* x_6; +x_6 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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: +{ +lean_object* x_6; +x_6 = l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__1___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5); +lean_dec(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("compiler", 8); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("enableNew", 9); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2; +x_3 = l_Lean_Name_mkStr2(x_1, x_2); +return x_3; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("(compiler) enable the new code generator, this should have no significant effect on your code but it does help to test the new code generator; unset to only use the old code generator instead", 191); +return x_1; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5() { +_start: +{ +uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = 0; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4; +x_4 = lean_box(x_1); +x_5 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_5, 0, x_4); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +return x_5; +} +} +static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____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_CoreM___hyg_6____closed__5; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2; +x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3694_(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3; +x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5; +x_4 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__6; +x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_compileDeclsNew___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = lean_lcnf_compile_decls(x_1, x_2, x_3, x_4); +return x_5; +} +} +LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_compileDecl___spec__1(lean_object* x_1, lean_object* x_2) { +_start: +{ +if (lean_obj_tag(x_1) == 0) +{ +lean_object* x_3; +x_3 = l_List_reverse___rarg(x_2); +return x_3; +} +else +{ +uint8_t x_4; +x_4 = !lean_is_exclusive(x_1); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_1, 0); +x_6 = lean_ctor_get(x_1, 1); +x_7 = l_Lean_MessageData_ofName(x_5); +lean_ctor_set(x_1, 1, x_2); +lean_ctor_set(x_1, 0, x_7); +{ +lean_object* _tmp_0 = x_6; +lean_object* _tmp_1 = x_1; +x_1 = _tmp_0; +x_2 = _tmp_1; +} +goto _start; +} +else +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_9 = lean_ctor_get(x_1, 0); +x_10 = lean_ctor_get(x_1, 1); +lean_inc(x_10); +lean_inc(x_9); +lean_dec(x_1); +x_11 = l_Lean_MessageData_ofName(x_9); +x_12 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_2); +x_1 = x_10; +x_2 = x_12; +goto _start; +} +} +} +} +static lean_object* _init_l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_inheritedTraceOptions; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1; +x_6 = lean_st_ref_get(x_5, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_6, 0); +x_9 = lean_ctor_get(x_2, 2); +x_10 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_8, x_9, x_1); +lean_dec(x_8); +x_11 = lean_box(x_10); +lean_ctor_set(x_6, 0, x_11); +return x_6; +} +else +{ +lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_12 = lean_ctor_get(x_6, 0); +x_13 = lean_ctor_get(x_6, 1); +lean_inc(x_13); +lean_inc(x_12); +lean_dec(x_6); +x_14 = lean_ctor_get(x_2, 2); +x_15 = l___private_Lean_Util_Trace_0__Lean_checkTraceOption(x_12, x_14, x_1); +lean_dec(x_12); +x_16 = lean_box(x_15); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_16); +lean_ctor_set(x_17, 1, x_13); +return x_17; +} } } -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg(lean_object* x_1, lean_object* x_2) { _start: { -uint8_t x_3; lean_object* x_4; -x_3 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__1(x_1, x_2); -x_4 = lean_box(x_3); -return x_4; -} -} -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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) { -_start: +lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; +x_3 = lean_st_ref_get(x_1, x_2); +x_4 = lean_ctor_get(x_3, 0); +lean_inc(x_4); +x_5 = lean_ctor_get(x_3, 1); +lean_inc(x_5); +lean_dec(x_3); +x_6 = lean_ctor_get(x_4, 3); +lean_inc(x_6); +lean_dec(x_4); +x_7 = lean_st_ref_take(x_1, x_5); +x_8 = lean_ctor_get(x_7, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_7, 1); +lean_inc(x_9); +lean_dec(x_7); +x_10 = !lean_is_exclusive(x_8); +if (x_10 == 0) { -lean_object* x_6; -x_6 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_5); -return x_6; -} -} -LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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: +lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_11 = lean_ctor_get(x_8, 3); +lean_dec(x_11); +x_12 = l_Lean_Core_State_traceState___default___closed__3; +lean_ctor_set(x_8, 3, x_12); +x_13 = lean_st_ref_set(x_1, x_8, x_9); +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_6; -x_6 = l_Lean_Declaration_foldExprM___at___private_Lean_CoreM_0__Lean_checkUnsupported___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_15; +x_15 = lean_ctor_get(x_13, 0); +lean_dec(x_15); +lean_ctor_set(x_13, 0, x_6); +return x_13; } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("compiler", 8); -return x_1; +lean_object* x_16; lean_object* x_17; +x_16 = lean_ctor_get(x_13, 1); +lean_inc(x_16); +lean_dec(x_13); +x_17 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_17, 0, x_6); +lean_ctor_set(x_17, 1, x_16); +return x_17; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("enableNew", 9); -return x_1; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_18 = lean_ctor_get(x_8, 0); +x_19 = lean_ctor_get(x_8, 1); +x_20 = lean_ctor_get(x_8, 2); +x_21 = lean_ctor_get(x_8, 4); +x_22 = lean_ctor_get(x_8, 5); +x_23 = lean_ctor_get(x_8, 6); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_dec(x_8); +x_24 = l_Lean_Core_State_traceState___default___closed__3; +x_25 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_25, 0, x_18); +lean_ctor_set(x_25, 1, x_19); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_24); +lean_ctor_set(x_25, 4, x_21); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_1, x_25, x_9); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +if (lean_is_scalar(x_28)) { + x_29 = lean_alloc_ctor(0, 2, 0); +} else { + x_29 = x_28; } +lean_ctor_set(x_29, 0, x_6); +lean_ctor_set(x_29, 1, x_27); +return x_29; } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3() { -_start: -{ -lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1; -x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2; -x_3 = l_Lean_Name_mkStr2(x_1, x_2); -return x_3; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4() { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4(lean_object* x_1) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("(compiler) enable the new code generator, this should have no significant effect on your code but it does help to test the new code generator; unset to only use the old code generator instead", 191); -return x_1; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg___boxed), 2, 0); +return x_2; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5() { +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -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_CoreM___hyg_3898____closed__1; -x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4; -x_4 = lean_box(x_1); -x_5 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_5, 0, x_4); -lean_ctor_set(x_5, 1, x_2); -lean_ctor_set(x_5, 2, x_3); -return x_5; -} +uint8_t x_8; +x_8 = !lean_is_exclusive(x_5); +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; lean_object* x_15; lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; lean_object* x_20; 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_9 = lean_ctor_get(x_5, 5); +x_10 = l_Lean_replaceRef(x_3, x_9); +lean_dec(x_9); +lean_ctor_set(x_5, 5, x_10); +x_11 = lean_st_ref_get(x_6, x_7); +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get(x_12, 3); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_PersistentArray_toArray___rarg(x_14); +x_16 = lean_array_get_size(x_15); +x_17 = lean_usize_of_nat(x_16); +lean_dec(x_16); +x_18 = 0; +x_19 = l_Array_mapMUnsafe_map___at___private_Lean_Util_Trace_0__Lean_addTraceNode___spec__1(x_17, x_18, x_15); +x_20 = lean_alloc_ctor(9, 3, 0); +lean_ctor_set(x_20, 0, x_2); +lean_ctor_set(x_20, 1, x_4); +lean_ctor_set(x_20, 2, x_19); +x_21 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_20, x_5, x_6, x_13); +lean_dec(x_5); +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_st_ref_take(x_6, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) +{ +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_25, 3); +lean_dec(x_28); +x_29 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_29, 0, x_3); +lean_ctor_set(x_29, 1, x_22); +x_30 = l_Lean_PersistentArray_push___rarg(x_1, x_29); +lean_ctor_set(x_25, 3, x_30); +x_31 = lean_st_ref_set(x_6, x_25, x_26); +x_32 = !lean_is_exclusive(x_31); +if (x_32 == 0) +{ +lean_object* x_33; lean_object* x_34; +x_33 = lean_ctor_get(x_31, 0); +lean_dec(x_33); +x_34 = lean_box(0); +lean_ctor_set(x_31, 0, x_34); +return x_31; } -static lean_object* _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6() { -_start: +else { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_initFn____x40_Lean_CoreM___hyg_6____closed__5; -x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1; -x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2; -x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); -return x_4; +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_31, 1); +lean_inc(x_35); +lean_dec(x_31); +x_36 = lean_box(0); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_35); +return x_37; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_CoreM___hyg_3898_(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3; -x_3 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5; -x_4 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6; -x_5 = l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(x_2, x_3, x_4, x_1); -return x_5; -} +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_38 = lean_ctor_get(x_25, 0); +x_39 = lean_ctor_get(x_25, 1); +x_40 = lean_ctor_get(x_25, 2); +x_41 = lean_ctor_get(x_25, 4); +x_42 = lean_ctor_get(x_25, 5); +x_43 = lean_ctor_get(x_25, 6); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_25); +x_44 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_44, 0, x_3); +lean_ctor_set(x_44, 1, x_22); +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_38); +lean_ctor_set(x_46, 1, x_39); +lean_ctor_set(x_46, 2, x_40); +lean_ctor_set(x_46, 3, x_45); +lean_ctor_set(x_46, 4, x_41); +lean_ctor_set(x_46, 5, x_42); +lean_ctor_set(x_46, 6, x_43); +x_47 = lean_st_ref_set(x_6, x_46, x_26); +x_48 = lean_ctor_get(x_47, 1); +lean_inc(x_48); +if (lean_is_exclusive(x_47)) { + lean_ctor_release(x_47, 0); + lean_ctor_release(x_47, 1); + x_49 = x_47; +} else { + lean_dec_ref(x_47); + x_49 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_compileDeclsNew___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = lean_lcnf_compile_decls(x_1, x_2, x_3, x_4); -return x_5; +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_List_mapTR_loop___at_Lean_compileDecl___spec__1(lean_object* x_1, lean_object* x_2) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_3; -x_3 = l_List_reverse___rarg(x_2); -return x_3; } else { -uint8_t x_4; -x_4 = !lean_is_exclusive(x_1); -if (x_4 == 0) -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_1, 0); -x_6 = lean_ctor_get(x_1, 1); -x_7 = l_Lean_MessageData_ofName(x_5); -lean_ctor_set(x_1, 1, x_2); -lean_ctor_set(x_1, 0, x_7); -{ -lean_object* _tmp_0 = x_6; -lean_object* _tmp_1 = x_1; -x_1 = _tmp_0; -x_2 = _tmp_1; +lean_object* x_52; 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; uint8_t 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_52 = lean_ctor_get(x_5, 0); +x_53 = lean_ctor_get(x_5, 1); +x_54 = lean_ctor_get(x_5, 2); +x_55 = lean_ctor_get(x_5, 3); +x_56 = lean_ctor_get(x_5, 4); +x_57 = lean_ctor_get(x_5, 5); +x_58 = lean_ctor_get(x_5, 6); +x_59 = lean_ctor_get(x_5, 7); +x_60 = lean_ctor_get(x_5, 8); +x_61 = lean_ctor_get(x_5, 9); +x_62 = lean_ctor_get(x_5, 10); +x_63 = lean_ctor_get_uint8(x_5, sizeof(void*)*11); +x_64 = lean_ctor_get_uint8(x_5, sizeof(void*)*11 + 1); +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_inc(x_52); +lean_dec(x_5); +x_65 = l_Lean_replaceRef(x_3, x_57); +lean_dec(x_57); +x_66 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_66, 0, x_52); +lean_ctor_set(x_66, 1, x_53); +lean_ctor_set(x_66, 2, x_54); +lean_ctor_set(x_66, 3, x_55); +lean_ctor_set(x_66, 4, x_56); +lean_ctor_set(x_66, 5, x_65); +lean_ctor_set(x_66, 6, x_58); +lean_ctor_set(x_66, 7, x_59); +lean_ctor_set(x_66, 8, x_60); +lean_ctor_set(x_66, 9, x_61); +lean_ctor_set(x_66, 10, x_62); +lean_ctor_set_uint8(x_66, sizeof(void*)*11, x_63); +lean_ctor_set_uint8(x_66, sizeof(void*)*11 + 1, x_64); +x_67 = lean_st_ref_get(x_6, x_7); +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, 0); +lean_ctor_set(x_76, 0, x_2); +lean_ctor_set(x_76, 1, x_4); +lean_ctor_set(x_76, 2, x_75); +x_77 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_76, x_66, x_6, 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_6, 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); } -goto _start; +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; } -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_1, 0); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_1); -x_11 = l_Lean_MessageData_ofName(x_9); -x_12 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_2); -x_1 = x_10; -x_2 = x_12; -goto _start; +lean_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_6, 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_compileDecl___spec__3(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_compileDecl___spec__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { if (lean_obj_tag(x_1) == 0) @@ -15449,21 +12985,39 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___ { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_inc(x_7); -x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_addDecl___spec__10(x_1, x_5, x_2, x_3, x_7, x_8, x_9); +x_10 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(x_1, x_5, x_2, x_3, x_7, x_8, x_9); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3(x_4, x_7, x_8, x_11); +x_12 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_4, x_7, x_8, x_11); lean_dec(x_7); lean_dec(x_4); return x_12; } } +static double _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; uint8_t x_2; double x_3; +x_1 = lean_unsigned_to_nat(0u); +x_2 = 0; +x_3 = l_Float_ofScientific(x_1, x_2, x_1); +return x_3; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_profiler; +return x_1; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___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, uint8_t x_8, double x_9, double x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { double x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1; +x_15 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__1; lean_inc(x_3); lean_inc(x_1); x_16 = lean_alloc_ctor(0, 2, 17); @@ -15472,7 +13026,7 @@ lean_ctor_set(x_16, 1, x_3); lean_ctor_set_float(x_16, sizeof(void*)*2, x_15); lean_ctor_set_float(x_16, sizeof(void*)*2 + 8, x_15); lean_ctor_set_uint8(x_16, sizeof(void*)*2 + 16, x_2); -x_17 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2; +x_17 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2; x_18 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_7, x_17); lean_dec(x_7); if (x_18 == 0) @@ -15517,6 +13071,23 @@ return x_26; } } } +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 53); +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___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, uint8_t x_7, double x_8, double x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { @@ -15546,25 +13117,71 @@ 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 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2; +x_21 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2; x_22 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2(x_1, x_2, x_3, x_4, x_15, x_5, x_6, x_7, x_8, x_9, x_21, x_12, x_13, x_20); lean_dec(x_13); return x_22; } } } +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_useHearbeats; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler; +return x_1; +} +} +static lean_object* _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_trace_profiler_threshold; +return x_1; +} +} +static double _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__4() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} +static double _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; double x_4; +x_1 = lean_unsigned_to_nat(1000000000u); +x_2 = 0; +x_3 = lean_unsigned_to_nat(0u); +x_4 = l_Float_ofScientific(x_1, x_2, x_3); +return x_4; +} +} LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___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; uint8_t x_16; lean_dec(x_8); -x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_addDecl___spec__9___rarg(x_10, x_11); +x_12 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___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); lean_inc(x_14); lean_dec(x_12); -x_15 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1; +x_15 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1; x_16 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_15); if (x_16 == 0) { @@ -15598,7 +13215,7 @@ x_95 = 0; x_96 = lean_unsigned_to_nat(0u); x_97 = l_Float_ofScientific(x_86, x_95, x_96); lean_dec(x_86); -x_98 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5; +x_98 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5; x_99 = lean_float_div(x_97, x_98); x_100 = l_Float_ofScientific(x_93, x_95, x_96); lean_dec(x_93); @@ -15635,7 +13252,7 @@ x_112 = 0; x_113 = lean_unsigned_to_nat(0u); x_114 = l_Float_ofScientific(x_86, x_112, x_113); lean_dec(x_86); -x_115 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5; +x_115 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5; x_116 = lean_float_div(x_114, x_115); x_117 = l_Float_ofScientific(x_110, x_112, x_113); lean_dec(x_110); @@ -15665,7 +13282,7 @@ lean_inc(x_21); x_22 = lean_ctor_get(x_19, 1); lean_inc(x_22); lean_dec(x_19); -x_70 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; +x_70 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2; x_71 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_70); if (x_71 == 0) { @@ -15680,13 +13297,13 @@ double x_73; double x_74; double x_75; lean_object* x_76; lean_object* x_77; uin x_73 = lean_unbox_float(x_22); x_74 = lean_unbox_float(x_21); x_75 = lean_float_sub(x_73, x_74); -x_76 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3; +x_76 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3; x_77 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_76); x_78 = 0; x_79 = lean_unsigned_to_nat(0u); x_80 = l_Float_ofScientific(x_77, x_78, x_79); lean_dec(x_77); -x_81 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4; +x_81 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__4; x_82 = lean_float_div(x_80, x_81); x_83 = lean_float_decLt(x_82, x_75); x_23 = x_83; @@ -15722,7 +13339,7 @@ x_30 = lean_st_ref_set(x_10, 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_compileDecl___spec__3(x_20, x_9, x_10, x_31); +x_32 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_20, x_9, x_10, x_31); lean_dec(x_10); lean_dec(x_9); lean_dec(x_20); @@ -15802,7 +13419,7 @@ x_50 = lean_st_ref_set(x_10, 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_compileDecl___spec__3(x_20, x_9, x_10, x_51); +x_52 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_20, x_9, x_10, x_51); lean_dec(x_10); lean_dec(x_9); lean_dec(x_20); @@ -15975,7 +13592,7 @@ lean_inc(x_127); x_128 = lean_ctor_get(x_125, 1); lean_inc(x_128); lean_dec(x_125); -x_176 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; +x_176 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2; x_177 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_1, x_176); if (x_177 == 0) { @@ -15990,7 +13607,7 @@ double x_179; double x_180; double x_181; lean_object* x_182; lean_object* x_183 x_179 = lean_unbox_float(x_128); x_180 = lean_unbox_float(x_127); x_181 = lean_float_sub(x_179, x_180); -x_182 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3; +x_182 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3; x_183 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_182); x_184 = 0; x_185 = lean_unsigned_to_nat(0u); @@ -16030,7 +13647,7 @@ x_136 = lean_st_ref_set(x_10, x_131, x_132); x_137 = lean_ctor_get(x_136, 1); lean_inc(x_137); lean_dec(x_136); -x_138 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3(x_126, x_9, x_10, x_137); +x_138 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_126, x_9, x_10, x_137); lean_dec(x_10); lean_dec(x_9); lean_dec(x_126); @@ -16110,7 +13727,7 @@ x_156 = lean_st_ref_set(x_10, x_155, x_132); x_157 = lean_ctor_get(x_156, 1); lean_inc(x_157); lean_dec(x_156); -x_158 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3(x_126, x_9, x_10, x_157); +x_158 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_126, x_9, x_10, x_157); lean_dec(x_10); lean_dec(x_9); lean_dec(x_126); @@ -16199,7 +13816,7 @@ lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_ctor_get(x_6, 2); lean_inc(x_9); lean_inc(x_1); -x_10 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_1, x_6, x_7, x_8); +x_10 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(x_1, x_6, x_7, x_8); x_11 = lean_ctor_get(x_10, 0); lean_inc(x_11); x_12 = lean_unbox(x_11); @@ -16209,7 +13826,7 @@ lean_object* x_13; lean_object* x_14; uint8_t x_15; x_13 = lean_ctor_get(x_10, 1); lean_inc(x_13); lean_dec(x_10); -x_14 = l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2; +x_14 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2; x_15 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_9, x_14); if (x_15 == 0) { @@ -16289,7 +13906,57 @@ return x_31; } } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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_Lean_throwError___at_Lean_compileDecl___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; uint8_t x_7; +x_5 = lean_ctor_get(x_2, 5); +x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); +x_7 = !lean_is_exclusive(x_6); +if (x_7 == 0) +{ +lean_object* x_8; lean_object* x_9; +x_8 = lean_ctor_get(x_6, 0); +lean_inc(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); +lean_ctor_set_tag(x_6, 1); +lean_ctor_set(x_6, 0, x_9); +return x_6; +} +else +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; +x_10 = lean_ctor_get(x_6, 0); +x_11 = lean_ctor_get(x_6, 1); +lean_inc(x_11); +lean_inc(x_10); +lean_dec(x_6); +lean_inc(x_5); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_5); +lean_ctor_set(x_12, 1, x_10); +x_13 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_13, 0, x_12); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +} +} +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_compileDecl___spec__7(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_ctor_get(x_2, 2); +lean_inc(x_5); +x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); +x_7 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_6, x_2, x_3, x_4); +lean_dec(x_2); +return x_7; +} +} +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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) { _start: { if (lean_obj_tag(x_3) == 0) @@ -16348,7 +14015,7 @@ x_40 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ 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 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_41, x_4, x_5, x_6); +x_42 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_41, x_4, x_5, x_6); x_43 = !lean_is_exclusive(x_42); if (x_43 == 0) { @@ -16417,7 +14084,7 @@ x_21 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ 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_throwError___at_Lean_addDecl___spec__2(x_22, x_4, x_5, x_10); +x_23 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_22, x_4, x_5, x_10); x_24 = !lean_is_exclusive(x_23); if (x_24 == 0) { @@ -16452,7 +14119,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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_List_foldlM___at_Lean_compileDecl___spec__12(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { if (lean_obj_tag(x_3) == 0) @@ -16511,7 +14178,7 @@ x_20 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ 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_Lean_addDecl___spec__2(x_21, x_4, x_5, x_6); +x_22 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_21, x_4, x_5, x_6); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { @@ -16544,7 +14211,7 @@ goto _start; } } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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: { if (lean_obj_tag(x_3) == 0) @@ -16579,7 +14246,7 @@ lean_inc(x_22); lean_dec(x_8); x_23 = lean_box(0); lean_inc(x_1); -x_24 = l_List_foldlM___at_Lean_compileDecl___spec__7(x_1, x_23, x_22, x_4, x_5, x_6); +x_24 = l_List_foldlM___at_Lean_compileDecl___spec__12(x_1, x_23, x_22, x_4, x_5, x_6); x_10 = x_24; goto block_18; } @@ -16605,7 +14272,7 @@ x_30 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ x_31 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_31, 0, x_29); lean_ctor_set(x_31, 1, x_30); -x_32 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_31, x_4, x_5, x_6); +x_32 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_31, x_4, x_5, x_6); x_33 = !lean_is_exclusive(x_32); if (x_33 == 0) { @@ -16636,7 +14303,7 @@ lean_inc(x_37); lean_dec(x_8); x_38 = lean_box(0); lean_inc(x_1); -x_39 = l_List_foldlM___at_Lean_compileDecl___spec__7(x_1, x_38, x_37, x_4, x_5, x_6); +x_39 = l_List_foldlM___at_Lean_compileDecl___spec__12(x_1, x_38, x_37, x_4, x_5, x_6); x_10 = x_39; goto block_18; } @@ -16684,7 +14351,7 @@ return x_17; } } } -LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___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: { switch (lean_obj_tag(x_2)) { @@ -16707,6 +14374,7 @@ x_11 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x if (lean_obj_tag(x_11) == 0) { lean_object* x_12; lean_object* x_13; +lean_dec(x_4); x_12 = lean_box(0); x_13 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_13, 0, x_12); @@ -16734,13 +14402,15 @@ x_19 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); -x_21 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_20, x_4, x_5, x_6); +x_21 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_20, x_4, x_5, x_6); +lean_dec(x_4); return x_21; } else { lean_object* x_22; lean_object* x_23; lean_dec(x_14); +lean_dec(x_4); x_22 = lean_box(0); x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); @@ -16796,7 +14466,8 @@ x_51 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_52, x_4, x_5, x_6); +x_53 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_52, x_4, x_5, x_6); +lean_dec(x_4); x_54 = !lean_is_exclusive(x_53); if (x_54 == 0) { @@ -16832,6 +14503,7 @@ x_30 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x if (lean_obj_tag(x_30) == 0) { lean_object* x_31; lean_object* x_32; +lean_dec(x_4); x_31 = lean_box(0); x_32 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_32, 0, x_31); @@ -16859,13 +14531,15 @@ x_38 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ 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_throwError___at_Lean_addDecl___spec__2(x_39, x_4, x_5, x_28); +x_40 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_39, x_4, x_5, x_28); +lean_dec(x_4); return x_40; } else { lean_object* x_41; lean_object* x_42; lean_dec(x_33); +lean_dec(x_4); x_41 = lean_box(0); x_42 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_42, 0, x_41); @@ -16922,7 +14596,8 @@ x_85 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ x_86 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_86, 0, x_84); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_86, x_4, x_5, x_6); +x_87 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_86, x_4, x_5, x_6); +lean_dec(x_4); x_88 = !lean_is_exclusive(x_87); if (x_88 == 0) { @@ -16958,6 +14633,7 @@ x_64 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x if (lean_obj_tag(x_64) == 0) { lean_object* x_65; lean_object* x_66; +lean_dec(x_4); x_65 = lean_box(0); x_66 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_66, 0, x_65); @@ -16985,13 +14661,15 @@ x_72 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__ x_73 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_73, 0, x_71); lean_ctor_set(x_73, 1, x_72); -x_74 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_73, x_4, x_5, x_62); +x_74 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_73, x_4, x_5, x_62); +lean_dec(x_4); return x_74; } else { lean_object* x_75; lean_object* x_76; lean_dec(x_67); +lean_dec(x_4); x_75 = lean_box(0); x_76 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_76, 0, x_75); @@ -17048,7 +14726,8 @@ x_119 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec_ x_120 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_120, 0, x_118); lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_120, x_4, x_5, x_6); +x_121 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_120, x_4, x_5, x_6); +lean_dec(x_4); x_122 = !lean_is_exclusive(x_121); if (x_122 == 0) { @@ -17084,6 +14763,7 @@ x_98 = l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x if (lean_obj_tag(x_98) == 0) { lean_object* x_99; lean_object* x_100; +lean_dec(x_4); x_99 = lean_box(0); x_100 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_100, 0, x_99); @@ -17111,13 +14791,15 @@ x_106 = l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec_ 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 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_107, x_4, x_5, x_96); +x_108 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_107, x_4, x_5, x_96); +lean_dec(x_4); return x_108; } else { lean_object* x_109; lean_object* x_110; lean_dec(x_101); +lean_dec(x_4); x_109 = lean_box(0); x_110 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_110, 0, x_109); @@ -17130,6 +14812,7 @@ return x_110; case 4: { lean_object* x_126; +lean_dec(x_4); lean_dec(x_1); x_126 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_126, 0, x_3); @@ -17142,7 +14825,8 @@ lean_object* x_127; lean_object* x_128; x_127 = lean_ctor_get(x_2, 0); lean_inc(x_127); lean_dec(x_2); -x_128 = l_List_foldlM___at_Lean_compileDecl___spec__6(x_1, x_3, x_127, x_4, x_5, x_6); +x_128 = l_List_foldlM___at_Lean_compileDecl___spec__11(x_1, x_3, x_127, x_4, x_5, x_6); +lean_dec(x_4); return x_128; } default: @@ -17151,13 +14835,14 @@ lean_object* x_129; lean_object* x_130; x_129 = lean_ctor_get(x_2, 2); lean_inc(x_129); lean_dec(x_2); -x_130 = l_List_foldlM___at_Lean_compileDecl___spec__8(x_1, x_3, x_129, x_4, x_5, x_6); +x_130 = l_List_foldlM___at_Lean_compileDecl___spec__13(x_1, x_3, x_129, x_4, x_5, x_6); +lean_dec(x_4); return x_130; } } } } -LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__9(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; @@ -17171,12 +14856,102 @@ x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); x_9 = lean_box(0); -x_10 = l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5(x_8, x_1, x_9, x_2, x_3, x_7); +x_10 = l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__10(x_8, x_1, x_9, x_2, x_3, x_7); lean_dec(x_3); -lean_dec(x_2); return x_10; } } +LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_compileDecl___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; uint8_t x_8; +x_5 = lean_st_ref_take(x_3, x_4); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); +lean_inc(x_7); +lean_dec(x_5); +x_8 = !lean_is_exclusive(x_6); +if (x_8 == 0) +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_9 = lean_ctor_get(x_6, 4); +lean_dec(x_9); +x_10 = lean_ctor_get(x_6, 0); +lean_dec(x_10); +x_11 = l_Lean_Core_instInhabitedCache___closed__1; +lean_ctor_set(x_6, 4, x_11); +lean_ctor_set(x_6, 0, x_1); +x_12 = lean_st_ref_set(x_3, x_6, x_7); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +x_15 = lean_box(0); +lean_ctor_set(x_12, 0, x_15); +return x_12; +} +else +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_12, 1); +lean_inc(x_16); +lean_dec(x_12); +x_17 = lean_box(0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_17); +lean_ctor_set(x_18, 1, x_16); +return x_18; +} +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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, 1); +x_20 = lean_ctor_get(x_6, 2); +x_21 = lean_ctor_get(x_6, 3); +x_22 = lean_ctor_get(x_6, 5); +x_23 = lean_ctor_get(x_6, 6); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_6); +x_24 = l_Lean_Core_instInhabitedCache___closed__1; +x_25 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_25, 0, x_1); +lean_ctor_set(x_25, 1, x_19); +lean_ctor_set(x_25, 2, x_20); +lean_ctor_set(x_25, 3, x_21); +lean_ctor_set(x_25, 4, x_24); +lean_ctor_set(x_25, 5, x_22); +lean_ctor_set(x_25, 6, x_23); +x_26 = lean_st_ref_set(x_3, x_25, x_7); +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +if (lean_is_exclusive(x_26)) { + lean_ctor_release(x_26, 0); + lean_ctor_release(x_26, 1); + x_28 = x_26; +} else { + lean_dec_ref(x_26); + x_28 = lean_box(0); +} +x_29 = lean_box(0); +if (lean_is_scalar(x_28)) { + x_30 = lean_alloc_ctor(0, 2, 0); +} else { + x_30 = x_28; +} +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_27); +return x_30; +} +} +} static lean_object* _init_l_Lean_compileDecl___lambda__1___closed__1() { _start: { @@ -17257,7 +15032,7 @@ static lean_object* _init_l_Lean_compileDecl___lambda__3___closed__1() { { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1; +x_2 = l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } @@ -17266,6 +15041,7 @@ LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__3(lean_object* x_1, lean_o _start: { 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_dec(x_4); x_8 = lean_alloc_closure((void*)(l_Lean_compileDecl___lambda__1___boxed), 5, 1); lean_closure_set(x_8, 0, x_1); lean_inc(x_3); @@ -17300,7 +15076,7 @@ lean_inc(x_17); lean_dec(x_15); lean_inc(x_6); lean_inc(x_5); -x_18 = l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__4(x_3, x_5, x_6, x_16); +x_18 = l___private_Lean_CoreM_0__Lean_checkUnsupported___at_Lean_compileDecl___spec__9(x_3, x_5, x_6, x_16); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; @@ -17311,7 +15087,7 @@ x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_17); x_21 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_21, 0, x_20); -x_22 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_21, x_5, x_6, x_19); +x_22 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_21, x_5, x_6, x_19); lean_dec(x_6); lean_dec(x_5); return x_22; @@ -17349,7 +15125,7 @@ lean_dec(x_3); x_27 = lean_ctor_get(x_13, 1); lean_inc(x_27); lean_dec(x_13); -x_28 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_15, x_5, x_6, x_27); +x_28 = l_Lean_throwKernelException___at_Lean_compileDecl___spec__7(x_15, x_5, x_6, x_27); lean_dec(x_6); return x_28; } @@ -17364,7 +15140,7 @@ lean_dec(x_13); x_30 = lean_ctor_get(x_14, 0); lean_inc(x_30); lean_dec(x_14); -x_31 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_30, x_5, x_6, x_29); +x_31 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_30, x_5, x_6, x_29); lean_dec(x_6); lean_dec(x_5); return x_31; @@ -17438,7 +15214,6 @@ x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = l_Lean_compileDecl___lambda__3(x_6, x_5, x_1, x_12, x_2, x_3, x_13); -lean_dec(x_12); return x_14; } else @@ -17471,11 +15246,48 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___rarg(x_1, x_2); +lean_dec(x_1); +return x_3; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_compileDecl___spec__4(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +_start: +{ +lean_object* x_8; +x_8 = l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_compileDecl___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +lean_dec(x_6); +return x_8; +} +} +LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { lean_object* x_5; -x_5 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__3(x_1, x_2, x_3, x_4); +x_5 = l_MonadExcept_ofExcept___at_Lean_compileDecl___spec__6(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -17547,46 +15359,74 @@ x_10 = l_Lean_withTraceNode___at_Lean_compileDecl___spec__2(x_1, x_2, x_3, x_9, return x_10; } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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_Lean_throwError___at_Lean_compileDecl___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} +LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_compileDecl___spec__7___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_throwKernelException___at_Lean_compileDecl___spec__7(x_1, x_2, x_3, x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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) { _start: { lean_object* x_7; -x_7 = l_List_foldlM___at_Lean_compileDecl___spec__6(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_List_foldlM___at_Lean_compileDecl___spec__11(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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_List_foldlM___at_Lean_compileDecl___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) { _start: { lean_object* x_7; -x_7 = l_List_foldlM___at_Lean_compileDecl___spec__7(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_List_foldlM___at_Lean_compileDecl___spec__12(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_compileDecl___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_List_foldlM___at_Lean_compileDecl___spec__8(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_List_foldlM___at_Lean_compileDecl___spec__13(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Declaration_foldExprM___at_Lean_compileDecl___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_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__5(x_1, x_2, x_3, x_4, x_5, x_6); +x_7 = l_Lean_Declaration_foldExprM___at_Lean_compileDecl___spec__10(x_1, x_2, x_3, x_4, x_5, x_6); lean_dec(x_5); -lean_dec(x_4); return x_7; } } +LEAN_EXPORT lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +lean_object* x_5; +x_5 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_1, x_2, x_3, x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_5; +} +} LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -17609,15 +15449,6 @@ lean_dec(x_1); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_compileDecl___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_Lean_compileDecl___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_4); -return x_8; -} -} LEAN_EXPORT lean_object* l_Lean_compileDecls___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { @@ -17648,14 +15479,14 @@ x_14 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_14, 0, x_13); x_15 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_15, 0, x_14); -x_16 = l_Lean_throwError___at_Lean_addDecl___spec__2(x_15, x_4, x_5, x_9); +x_16 = l_Lean_throwError___at_Lean_compileDecl___spec__8(x_15, x_4, x_5, x_9); lean_dec(x_4); return x_16; } else { lean_object* x_17; -x_17 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_12, x_4, x_5, x_9); +x_17 = l_Lean_throwKernelException___at_Lean_compileDecl___spec__7(x_12, x_4, x_5, x_9); return x_17; } } @@ -17665,7 +15496,7 @@ lean_object* x_18; lean_object* x_19; x_18 = lean_ctor_get(x_11, 0); lean_inc(x_18); lean_dec(x_11); -x_19 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_18, x_4, x_5, x_9); +x_19 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_18, x_4, x_5, x_9); lean_dec(x_4); return x_19; } @@ -17752,50 +15583,6 @@ lean_dec(x_1); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_addAndCompile(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -lean_inc(x_3); -lean_inc(x_2); -lean_inc(x_1); -x_5 = l_Lean_addDecl(x_1, x_2, x_3, x_4); -if (lean_obj_tag(x_5) == 0) -{ -lean_object* x_6; lean_object* x_7; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = l_Lean_compileDecl(x_1, x_2, x_3, x_6); -return x_7; -} -else -{ -uint8_t x_8; -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_8 = !lean_is_exclusive(x_5); -if (x_8 == 0) -{ -return x_5; -} -else -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); -lean_inc(x_9); -lean_dec(x_5); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_9); -lean_ctor_set(x_11, 1, x_10); -return x_11; -} -} -} -} LEAN_EXPORT uint8_t l_Lean_getDiag(lean_object* x_1) { _start: { @@ -19561,39 +17348,6 @@ l_Lean_mkArrow___closed__1 = _init_l_Lean_mkArrow___closed__1(); lean_mark_persistent(l_Lean_mkArrow___closed__1); l_Lean_mkArrow___closed__2 = _init_l_Lean_mkArrow___closed__2(); lean_mark_persistent(l_Lean_mkArrow___closed__2); -l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1 = _init_l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1(); -lean_mark_persistent(l_Lean_logAt___at_Lean_addDecl___spec__6___closed__1); -l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1(); -lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8___closed__1); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__1(); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__2___closed__2); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__1); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__3___closed__2); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__1); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__2); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__3); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__4(); -l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5 = _init_l_Lean_withTraceNode___at_Lean_addDecl___spec__7___lambda__4___closed__5(); -l_Lean_addDecl___lambda__1___closed__1 = _init_l_Lean_addDecl___lambda__1___closed__1(); -lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__1); -l_Lean_addDecl___lambda__1___closed__2 = _init_l_Lean_addDecl___lambda__1___closed__2(); -lean_mark_persistent(l_Lean_addDecl___lambda__1___closed__2); -l_Lean_addDecl___lambda__3___closed__1 = _init_l_Lean_addDecl___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__1); -l_Lean_addDecl___lambda__3___closed__2 = _init_l_Lean_addDecl___lambda__3___closed__2(); -lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__2); -l_Lean_addDecl___lambda__3___closed__3 = _init_l_Lean_addDecl___lambda__3___closed__3(); -lean_mark_persistent(l_Lean_addDecl___lambda__3___closed__3); -l_Lean_addDecl___closed__1 = _init_l_Lean_addDecl___closed__1(); -lean_mark_persistent(l_Lean_addDecl___closed__1); -l_Lean_addDecl___closed__2 = _init_l_Lean_addDecl___closed__2(); -lean_mark_persistent(l_Lean_addDecl___closed__2); l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1 = _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1(); lean_mark_persistent(l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__1); l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2 = _init_l___private_Lean_CoreM_0__Lean_supportedRecursors___closed__2(); @@ -19662,24 +17416,41 @@ l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rar lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__3); l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4 = _init_l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4(); lean_mark_persistent(l_List_foldlM___at___private_Lean_CoreM_0__Lean_checkUnsupported___spec__2___rarg___lambda__2___closed__4); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__1); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__2); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__3); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__4); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__5); -l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3898____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_3898_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__1); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__2); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__3); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__4); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__5); +l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__6 = _init_l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_CoreM___hyg_3694____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_CoreM___hyg_3694_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_compiler_enableNew = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_compiler_enableNew); lean_dec_ref(res); -}l_Lean_compileDecl___lambda__1___closed__1 = _init_l_Lean_compileDecl___lambda__1___closed__1(); +}l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1 = _init_l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3___closed__1); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__1 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__1(); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__2___closed__2); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__1); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__3___closed__2); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__1); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__2); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3(); +lean_mark_persistent(l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__3); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__4 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__4(); +l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5 = _init_l_Lean_withTraceNode___at_Lean_compileDecl___spec__2___lambda__4___closed__5(); +l_Lean_compileDecl___lambda__1___closed__1 = _init_l_Lean_compileDecl___lambda__1___closed__1(); lean_mark_persistent(l_Lean_compileDecl___lambda__1___closed__1); l_Lean_compileDecl___lambda__1___closed__2 = _init_l_Lean_compileDecl___lambda__1___closed__2(); lean_mark_persistent(l_Lean_compileDecl___lambda__1___closed__2); diff --git a/stage0/stdlib/Lean/Data/Options.c b/stage0/stdlib/Lean/Data/Options.c index bd2ebf4b44e2..1905382bfb99 100644 --- a/stage0/stdlib/Lean/Data/Options.c +++ b/stage0/stdlib/Lean/Data/Options.c @@ -56,6 +56,7 @@ static lean_object* l_Lean_Option_registerBuiltinOption___closed__15; static lean_object* l_Lean_Option_registerBuiltinOption___closed__8; static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__24; LEAN_EXPORT lean_object* l_Lean_Option_set(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Data_Options___hyg_1050_; LEAN_EXPORT lean_object* l_Lean_Option_get___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instToStringOptions; LEAN_EXPORT lean_object* l_Lean_registerOption___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -138,7 +139,6 @@ lean_object* l_Array_append___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__33; static lean_object* l_Lean_Option_registerOption___closed__11; LEAN_EXPORT lean_object* l_Lean_withInPattern___rarg___lambda__1(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Data_Options___hyg_1080_; LEAN_EXPORT lean_object* l_Lean_Option_register(lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__35; lean_object* l_Lean_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -170,6 +170,7 @@ LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___lambda__1(lean_object*, l lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__5; LEAN_EXPORT lean_object* l_Lean_getNatOption___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_143_(lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_KVMap_getBool(lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_instForInOptionsProdNameDataValue___closed__1; @@ -216,7 +217,6 @@ static lean_object* l___auto____x40_Lean_Data_Options___hyg_69____closed__3; static lean_object* l_Lean_Option_registerBuiltinOption___closed__28; lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_registerOption___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_173_(lean_object*); static lean_object* l_Lean_Option___aux__Lean__Data__Options______macroRules__Lean__Option__registerBuiltinOption__1___closed__15; static lean_object* l___auto____x40_Lean_Data_Options___hyg_69____closed__4; LEAN_EXPORT lean_object* l_Lean_getBoolOption___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -676,7 +676,7 @@ x_1 = lean_box(0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_173_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Data_Options___hyg_143_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1688,7 +1688,7 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Option_setIfNotSet___rarg), 4, 0); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Data_Options___hyg_1080_() { +static lean_object* _init_l___auto____x40_Lean_Data_Options___hyg_1050_() { _start: { lean_object* x_1; @@ -3216,7 +3216,7 @@ l_Lean_instInhabitedOptionDecl = _init_l_Lean_instInhabitedOptionDecl(); lean_mark_persistent(l_Lean_instInhabitedOptionDecl); l_Lean_instInhabitedOptionDecls = _init_l_Lean_instInhabitedOptionDecls(); lean_mark_persistent(l_Lean_instInhabitedOptionDecls); -if (builtin) {res = l_Lean_initFn____x40_Lean_Data_Options___hyg_173_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Data_Options___hyg_143_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Data_Options_0__Lean_optionDeclsRef = lean_io_result_get_value(res); lean_mark_persistent(l___private_Lean_Data_Options_0__Lean_optionDeclsRef); @@ -3245,8 +3245,8 @@ l_Lean_Option_Decl_group___default = _init_l_Lean_Option_Decl_group___default(); lean_mark_persistent(l_Lean_Option_Decl_group___default); l_Lean_Option_Decl_descr___default = _init_l_Lean_Option_Decl_descr___default(); lean_mark_persistent(l_Lean_Option_Decl_descr___default); -l___auto____x40_Lean_Data_Options___hyg_1080_ = _init_l___auto____x40_Lean_Data_Options___hyg_1080_(); -lean_mark_persistent(l___auto____x40_Lean_Data_Options___hyg_1080_); +l___auto____x40_Lean_Data_Options___hyg_1050_ = _init_l___auto____x40_Lean_Data_Options___hyg_1050_(); +lean_mark_persistent(l___auto____x40_Lean_Data_Options___hyg_1050_); l_Lean_Option_registerBuiltinOption___closed__1 = _init_l_Lean_Option_registerBuiltinOption___closed__1(); lean_mark_persistent(l_Lean_Option_registerBuiltinOption___closed__1); l_Lean_Option_registerBuiltinOption___closed__2 = _init_l_Lean_Option_registerBuiltinOption___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/App.c b/stage0/stdlib/Lean/Elab/App.c index 721ac54cd93e..32c3dceea9f3 100644 --- a/stage0/stdlib/Lean/Elab/App.c +++ b/stage0/stdlib/Lean/Elab/App.c @@ -4553,7 +4553,7 @@ lean_object* x_11; uint8_t x_12; uint8_t x_13; lean_object* x_14; x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = 1; +x_12 = 0; x_13 = 0; lean_inc(x_8); lean_inc(x_7); @@ -45853,7 +45853,7 @@ x_65 = l_Lean_Elab_Term_SavedState_restore(x_44, x_64, x_5, x_6, x_7, x_8, x_9, x_66 = lean_ctor_get(x_65, 1); lean_inc(x_66); lean_dec(x_65); -x_67 = 1; +x_67 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); diff --git a/stage0/stdlib/Lean/Elab/Binders.c b/stage0/stdlib/Lean/Elab/Binders.c index 81d6dca81ce4..8d3071393a60 100644 --- a/stage0/stdlib/Lean/Elab/Binders.c +++ b/stage0/stdlib/Lean/Elab/Binders.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Elab.Binders -// Imports: Lean.Elab.Quotation.Precheck Lean.Elab.Term Lean.Elab.BindersUtil Lean.Elab.PreDefinition.WF.TerminationHint +// Imports: Lean.Elab.Quotation.Precheck Lean.Elab.Term Lean.Elab.BindersUtil Lean.Elab.SyntheticMVars Lean.Elab.PreDefinition.WF.TerminationHint #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -64,6 +64,8 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__20; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDepArrow(lean_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_elabLetDeclCore(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__5; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5; static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__3; @@ -76,6 +78,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__6(lean_object static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___closed__2; static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__11; lean_object* l_Lean_mkAppB(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange___closed__2; @@ -91,14 +94,12 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange(lean static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_FunBinders_elabFunBindersAux(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__3; static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_precheckFun___spec__4___closed__1; uint8_t l_Lean_Syntax_isAntiquotSuffixSplice(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2; static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__16; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__8; @@ -135,10 +136,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declRange__ lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__3; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl(lean_object*); -lean_object* l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__7; lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6; static lean_object* l_Lean_Elab_Term_elabArrow___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclCore___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__4; @@ -179,7 +180,6 @@ lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckFun(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*); @@ -192,7 +192,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandB static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__1; static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__3; lean_object* l_Lean_mkApp4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun___closed__3; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_expandFun___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); @@ -216,7 +215,6 @@ LEAN_EXPORT lean_object* l_Lean_mkFreshId___at___private_Lean_Elab_Binders_0__Le static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBindersAux_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBindersEx___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews(lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___closed__2; @@ -231,6 +229,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed_ static lean_object* l_Lean_Elab_Term_precheckArrow___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__27; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow___closed__1; @@ -243,7 +242,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRange___ static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__14; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_elabLetDeclCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at_Lean_Elab_Term_expandFunBinders_loop___spec__2(lean_object*, lean_object*); @@ -364,6 +362,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange(le static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__3; lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___closed__5; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4; lean_object* l_List_forM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun___closed__1; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__17; @@ -386,6 +385,7 @@ lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContex lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_precheckFun___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_Term_initFn____x40_Lean_Elab_Binders___hyg_10163_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFunBinders(lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -428,7 +428,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsIntoMatch(lean_object*, LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun(lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__2; uint8_t lean_name_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3; static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__12; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Term_kindOfBinderName(lean_object*); @@ -443,9 +442,9 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFail LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_elabBinders___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__4; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__3; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -456,23 +455,24 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_ 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*); uint8_t l_Lean_Syntax_isIdent(lean_object*); lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_elabFun___spec__5___rarg(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_precheckFun___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop___rarg___closed__1; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_precheckFun___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandOptIdent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__5; +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__4; +lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(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_Term_quoteAutoTactic___spec__6___closed__6; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__7; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__1; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__5; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__2; lean_object* l_Lean_LocalContext_mkLocalDecl(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); @@ -504,6 +504,7 @@ static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__5; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandWhereDecls___lambda__1(lean_object*); static lean_object* l_Lean_Elab_Term_expandForall___closed__3; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8; lean_object* l_Lean_Elab_Term_Quotation_precheck(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_instToExprPreresolved___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__3; @@ -541,7 +542,6 @@ static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderView LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_propagateExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandFunBinders_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange___closed__6; uint8_t l_Lean_Name_isAtomic(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5___rarg(lean_object*); @@ -563,7 +563,6 @@ static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787_ lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_elabBinderViews_loop(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10; lean_object* lean_environment_main_module(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_quoteAutoTactic___spec__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_declareTacticSyntax___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*); @@ -587,9 +586,11 @@ static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14; lean_object* l_Lean_mkSepArray(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___closed__2; lean_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_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13; static lean_object* l_Lean_Elab_Term_elabArrow___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_registerFailedToInferBinderTypeInfo___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13___rarg(lean_object*); @@ -607,6 +608,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___close LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinder(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_expandForall_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Elab_Term_precheckFun___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___boxed(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_Term_quoteAutoTactic___spec__6___closed__19; static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabLetDeclAux___closed__4; @@ -621,7 +623,7 @@ lean_object* l_Lean_Elab_Term_clearInMatch(lean_object*, lean_object*, lean_obje static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_expandMatchAltsWhereDecls___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__7; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11; static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__26; @@ -641,7 +643,6 @@ static lean_object* l_Lean_Elab_Term_expandForall___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___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___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__4; lean_object* l___private_Lean_ToExpr_0__Lean_Name_toExprAux(lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__20; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__4; @@ -654,7 +655,6 @@ static lean_object* l_Lean_Elab_Term_expandWhereDecls___closed__9; static lean_object* l_Lean_Elab_Term_expandLetEqnsDecl___closed__2; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabFun___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__10; @@ -680,6 +680,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange(lean_ static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetFunDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandMatchAltsIntoMatchAux___closed__8; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_docString___closed__1; static lean_object* l_Lean_Elab_Term_expandForall___closed__4; @@ -687,7 +688,6 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRan static lean_object* l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___closed__4; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__21; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_checkLocalInstanceParameters___closed__4; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderType(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandSimpleBinderWithType___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__6; @@ -708,7 +708,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabFun_declRange___closed__6; lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Syntax_SepArray_getElems___spec__1(lean_object*, size_t, size_t, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___closed__3; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14; uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabLetDeclCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg(lean_object*, lean_object*, lean_object*); @@ -727,6 +726,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabBinders___rarg___lambda__1(lean_ob LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshBinderName___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___spec__2___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ToExpr_0__Lean_List_toExprAux___at_Lean_Elab_Term_quoteAutoTactic___spec__6___closed__17; static lean_object* l___regBuiltin_Lean_Elab_Term_precheckArrow___closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15; static lean_object* l_Lean_Elab_Term_expandFunBinders_loop___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___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_throwUnsupportedSyntax___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_toBinderViews___spec__1___rarg___closed__2; @@ -757,6 +757,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForall_declRange___closed_ lean_object* l_Lean_Elab_Term_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange___closed__3; static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__15; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3; static lean_object* l_Lean_Elab_Term_declareTacticSyntax___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_precheckFun___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_quoteAutoTactic___closed__9; @@ -772,7 +773,6 @@ uint8_t l_Lean_Syntax_isAntiquotSplice(lean_object*); static lean_object* l_Lean_Elab_Term_declareTacticSyntax___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandFun_declRange___closed__7; static lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderModifier___closed__2; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5; lean_object* l_Lean_Elab_Term_elabType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___at_Lean_Elab_Term_expandFunBinders_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabArrow(lean_object*); @@ -4345,72 +4345,6 @@ x_13 = l_Lean_Elab_Term_addTermInfo_x27(x_1, x_2, x_10, x_10, x_11, x_12, x_3, x return x_13; } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -uint8_t x_10; -x_10 = !lean_is_exclusive(x_7); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_11 = lean_ctor_get(x_7, 5); -x_12 = l_Lean_replaceRef(x_1, x_11); -lean_dec(x_11); -lean_ctor_set(x_7, 5, x_12); -x_13 = l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_7); -return x_13; -} -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_14 = lean_ctor_get(x_7, 0); -x_15 = lean_ctor_get(x_7, 1); -x_16 = lean_ctor_get(x_7, 2); -x_17 = lean_ctor_get(x_7, 3); -x_18 = lean_ctor_get(x_7, 4); -x_19 = lean_ctor_get(x_7, 5); -x_20 = lean_ctor_get(x_7, 6); -x_21 = lean_ctor_get(x_7, 7); -x_22 = lean_ctor_get(x_7, 8); -x_23 = lean_ctor_get(x_7, 9); -x_24 = lean_ctor_get(x_7, 10); -x_25 = lean_ctor_get_uint8(x_7, sizeof(void*)*11); -x_26 = lean_ctor_get_uint8(x_7, sizeof(void*)*11 + 1); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_20); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_dec(x_7); -x_27 = l_Lean_replaceRef(x_1, x_19); -lean_dec(x_19); -x_28 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_28, 0, x_14); -lean_ctor_set(x_28, 1, x_15); -lean_ctor_set(x_28, 2, x_16); -lean_ctor_set(x_28, 3, x_17); -lean_ctor_set(x_28, 4, x_18); -lean_ctor_set(x_28, 5, x_27); -lean_ctor_set(x_28, 6, x_20); -lean_ctor_set(x_28, 7, x_21); -lean_ctor_set(x_28, 8, x_22); -lean_ctor_set(x_28, 9, x_23); -lean_ctor_set(x_28, 10, x_24); -lean_ctor_set_uint8(x_28, sizeof(void*)*11, x_25); -lean_ctor_set_uint8(x_28, sizeof(void*)*11 + 1, x_26); -x_29 = l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(x_2, x_3, x_4, x_5, x_6, x_28, x_8, x_9); -lean_dec(x_28); -return x_29; -} -} -} static lean_object* _init_l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___closed__1() { _start: { @@ -4450,8 +4384,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureA { lean_object* x_9; lean_object* x_10; lean_object* x_11; uint8_t x_12; x_9 = lean_ctor_get(x_1, 1); -lean_inc(x_9); -lean_dec(x_1); x_10 = l_Lean_Syntax_getId(x_9); x_11 = lean_erase_macro_scopes(x_10); x_12 = l_Lean_Name_isAtomic(x_11); @@ -4467,24 +4399,14 @@ x_16 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___ x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); -x_18 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_9, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_9); +x_18 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_9, x_17, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_18; } else { lean_object* x_19; lean_object* x_20; lean_dec(x_11); -lean_dec(x_9); -lean_dec(x_7); lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); lean_dec(x_2); x_19 = lean_box(0); x_20 = lean_alloc_ctor(0, 2, 0); @@ -4494,17 +4416,17 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___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___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_10; -x_10 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_9; +x_9 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); +lean_dec(x_3); lean_dec(x_1); -return x_10; +return x_9; } } static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_1787____closed__1() { @@ -5107,13 +5029,8 @@ else { lean_object* x_15; lean_object* x_16; x_15 = lean_array_fget(x_1, x_3); -lean_inc(x_10); lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); lean_inc(x_5); -lean_inc(x_15); x_16 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName(x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_16) == 0) { @@ -6238,7 +6155,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(264u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6250,7 +6167,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(267u); +x_1 = lean_unsigned_to_nat(268u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6278,7 +6195,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(264u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6290,7 +6207,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandForall_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(263u); +x_1 = lean_unsigned_to_nat(264u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6545,7 +6462,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6557,7 +6474,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(275u); +x_1 = lean_unsigned_to_nat(276u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6585,7 +6502,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -6597,7 +6514,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabForall_declRange___c _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(269u); +x_1 = lean_unsigned_to_nat(270u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7019,7 +6936,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7031,7 +6948,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(291u); +x_1 = lean_unsigned_to_nat(292u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7059,7 +6976,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7071,7 +6988,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabArrow_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(284u); +x_1 = lean_unsigned_to_nat(285u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7225,7 +7142,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(297u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7237,7 +7154,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(302u); +x_1 = lean_unsigned_to_nat(303u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7265,7 +7182,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(297u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -7277,7 +7194,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabDepArrow_declRange__ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(297u); +x_1 = lean_unsigned_to_nat(298u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -14869,13 +14786,8 @@ else { lean_object* x_14; lean_object* x_15; x_14 = lean_array_fget(x_1, x_2); -lean_inc(x_9); lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); lean_inc(x_4); -lean_inc(x_14); x_15 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName(x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_15) == 0) { @@ -18034,7 +17946,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(595u); +x_1 = lean_unsigned_to_nat(596u); x_2 = lean_unsigned_to_nat(41u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18046,7 +17958,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(606u); +x_1 = lean_unsigned_to_nat(607u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18074,7 +17986,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(595u); +x_1 = lean_unsigned_to_nat(596u); x_2 = lean_unsigned_to_nat(45u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18086,7 +17998,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandFun_declRange___cl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(595u); +x_1 = lean_unsigned_to_nat(596u); x_2 = lean_unsigned_to_nat(54u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18245,7 +18157,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(608u); +x_1 = lean_unsigned_to_nat(609u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18257,7 +18169,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(611u); +x_1 = lean_unsigned_to_nat(612u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18285,7 +18197,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(608u); +x_1 = lean_unsigned_to_nat(609u); x_2 = lean_unsigned_to_nat(50u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -18297,7 +18209,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_expandExplicitFun_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(608u); +x_1 = lean_unsigned_to_nat(609u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20282,7 +20194,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(625u); +x_1 = lean_unsigned_to_nat(626u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20294,7 +20206,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(638u); +x_1 = lean_unsigned_to_nat(639u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20322,7 +20234,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(625u); +x_1 = lean_unsigned_to_nat(626u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20334,7 +20246,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabFun_declRange___clos _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(625u); +x_1 = lean_unsigned_to_nat(626u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -20670,133 +20582,136 @@ x_12 = l_Array_unzip___rarg(x_4); 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_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; x_14 = lean_ctor_get(x_12, 0); x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_1); +x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); +lean_closure_set(x_16, 0, x_1); +x_17 = 2; 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_1); -x_16 = l_Lean_Elab_Term_elabType(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_16) == 0) +x_18 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_16, x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_16, 1); -lean_inc(x_18); -lean_dec(x_16); -x_19 = l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__3; -x_20 = l_Lean_Elab_Term_registerCustomErrorIfMVar(x_17, x_1, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__3; +x_22 = l_Lean_Elab_Term_registerCustomErrorIfMVar(x_19, x_1, x_21, x_5, x_6, x_7, x_8, x_9, x_10, x_20); if (x_2 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -lean_inc(x_17); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_17); -x_23 = lean_box(0); -x_24 = 1; +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +lean_inc(x_19); +x_24 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_24, 0, x_19); +x_25 = lean_box(0); +x_26 = 1; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_25 = l_Lean_Elab_Term_elabTermEnsuringType(x_3, x_22, x_24, x_24, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -if (lean_obj_tag(x_25) == 0) +x_27 = l_Lean_Elab_Term_elabTermEnsuringType(x_3, x_24, x_26, x_26, x_25, x_5, x_6, x_7, x_8, x_9, x_10, x_23); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_26; lean_object* x_27; uint8_t x_28; uint8_t x_29; lean_object* x_30; -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 = 0; -x_29 = 1; +lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; +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 = 0; +x_31 = 1; lean_inc(x_15); -x_30 = l_Lean_Meta_mkForallFVars(x_15, x_17, x_28, x_24, x_29, x_7, x_8, x_9, x_10, x_27); -if (lean_obj_tag(x_30) == 0) +x_32 = l_Lean_Meta_mkForallFVars(x_15, x_19, x_30, x_26, x_31, x_7, x_8, x_9, x_10, x_29); +if (lean_obj_tag(x_32) == 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_mkLambdaFVars(x_15, x_26, x_28, x_28, x_29, x_7, x_8, x_9, x_10, x_32); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_32, 0); +lean_inc(x_33); +x_34 = lean_ctor_get(x_32, 1); +lean_inc(x_34); +lean_dec(x_32); +x_35 = l_Lean_Meta_mkLambdaFVars(x_15, x_28, x_30, x_30, x_31, x_7, x_8, x_9, x_10, x_34); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -if (lean_obj_tag(x_33) == 0) +if (lean_obj_tag(x_35) == 0) { -uint8_t x_34; -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_33, 0); +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_35, 0); lean_ctor_set(x_12, 1, x_14); -lean_ctor_set(x_12, 0, x_35); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_31); -lean_ctor_set(x_36, 1, x_12); -lean_ctor_set(x_33, 0, x_36); -return x_33; +lean_ctor_set(x_12, 0, x_37); +x_38 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_38, 0, x_33); +lean_ctor_set(x_38, 1, x_12); +lean_ctor_set(x_35, 0, x_38); +return x_35; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_33, 0); -x_38 = lean_ctor_get(x_33, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_33); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_39 = lean_ctor_get(x_35, 0); +x_40 = lean_ctor_get(x_35, 1); +lean_inc(x_40); +lean_inc(x_39); +lean_dec(x_35); lean_ctor_set(x_12, 1, x_14); -lean_ctor_set(x_12, 0, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_31); -lean_ctor_set(x_39, 1, x_12); -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; +lean_ctor_set(x_12, 0, x_39); +x_41 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_41, 0, x_33); +lean_ctor_set(x_41, 1, x_12); +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_41; -lean_dec(x_31); +uint8_t x_43; +lean_dec(x_33); lean_free_object(x_12); lean_dec(x_14); -x_41 = !lean_is_exclusive(x_33); -if (x_41 == 0) +x_43 = !lean_is_exclusive(x_35); +if (x_43 == 0) { -return x_33; +return x_35; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_33, 0); -x_43 = lean_ctor_get(x_33, 1); -lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_33); -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_35, 0); +x_45 = lean_ctor_get(x_35, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_35); +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_26); +uint8_t x_47; +lean_dec(x_28); lean_free_object(x_12); lean_dec(x_15); lean_dec(x_14); @@ -20804,30 +20719,30 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_45 = !lean_is_exclusive(x_30); -if (x_45 == 0) +x_47 = !lean_is_exclusive(x_32); +if (x_47 == 0) { -return x_30; +return x_32; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_30, 0); -x_47 = lean_ctor_get(x_30, 1); -lean_inc(x_47); -lean_inc(x_46); -lean_dec(x_30); -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_32, 0); +x_49 = lean_ctor_get(x_32, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_32); +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_49; -lean_dec(x_17); +uint8_t x_51; +lean_dec(x_19); lean_free_object(x_12); lean_dec(x_15); lean_dec(x_14); @@ -20835,121 +20750,121 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_49 = !lean_is_exclusive(x_25); -if (x_49 == 0) +x_51 = !lean_is_exclusive(x_27); +if (x_51 == 0) { -return x_25; +return x_27; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_25, 0); -x_51 = lean_ctor_get(x_25, 1); -lean_inc(x_51); -lean_inc(x_50); -lean_dec(x_25); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_27, 0); +x_53 = lean_ctor_get(x_27, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_27); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } } else { -lean_object* x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; lean_object* x_57; +lean_object* x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); -x_53 = lean_ctor_get(x_20, 1); -lean_inc(x_53); -lean_dec(x_20); -x_54 = 0; -x_55 = 1; -x_56 = 1; -x_57 = l_Lean_Meta_mkForallFVars(x_15, x_17, x_54, x_55, x_56, x_7, x_8, x_9, x_10, x_53); -if (lean_obj_tag(x_57) == 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; uint8_t x_64; -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); -lean_inc(x_58); -x_60 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_60, 0, x_58); -x_61 = 0; -x_62 = lean_box(0); -x_63 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_60, x_61, x_62, x_7, x_8, x_9, x_10, x_59); +x_55 = lean_ctor_get(x_22, 1); +lean_inc(x_55); +lean_dec(x_22); +x_56 = 0; +x_57 = 1; +x_58 = 1; +x_59 = l_Lean_Meta_mkForallFVars(x_15, x_19, x_56, x_57, x_58, x_7, x_8, x_9, x_10, x_55); +if (lean_obj_tag(x_59) == 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; uint8_t x_66; +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); +lean_inc(x_60); +x_62 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_62, 0, x_60); +x_63 = 0; +x_64 = lean_box(0); +x_65 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_62, x_63, x_64, x_7, x_8, x_9, x_10, x_61); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +x_66 = !lean_is_exclusive(x_65); +if (x_66 == 0) { -lean_object* x_65; lean_object* x_66; -x_65 = lean_ctor_get(x_63, 0); +lean_object* x_67; lean_object* x_68; +x_67 = lean_ctor_get(x_65, 0); lean_ctor_set(x_12, 1, x_14); -lean_ctor_set(x_12, 0, x_65); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_58); -lean_ctor_set(x_66, 1, x_12); -lean_ctor_set(x_63, 0, x_66); -return x_63; +lean_ctor_set(x_12, 0, x_67); +x_68 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_68, 0, x_60); +lean_ctor_set(x_68, 1, x_12); +lean_ctor_set(x_65, 0, x_68); +return x_65; } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_67 = lean_ctor_get(x_63, 0); -x_68 = lean_ctor_get(x_63, 1); -lean_inc(x_68); -lean_inc(x_67); -lean_dec(x_63); +lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +x_69 = lean_ctor_get(x_65, 0); +x_70 = lean_ctor_get(x_65, 1); +lean_inc(x_70); +lean_inc(x_69); +lean_dec(x_65); lean_ctor_set(x_12, 1, x_14); -lean_ctor_set(x_12, 0, x_67); -x_69 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_69, 0, x_58); -lean_ctor_set(x_69, 1, x_12); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_69); -lean_ctor_set(x_70, 1, x_68); -return x_70; +lean_ctor_set(x_12, 0, x_69); +x_71 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_71, 0, x_60); +lean_ctor_set(x_71, 1, x_12); +x_72 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_72, 0, x_71); +lean_ctor_set(x_72, 1, x_70); +return x_72; } } else { -uint8_t x_71; +uint8_t x_73; lean_free_object(x_12); lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_71 = !lean_is_exclusive(x_57); -if (x_71 == 0) +x_73 = !lean_is_exclusive(x_59); +if (x_73 == 0) { -return x_57; +return x_59; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_57, 0); -x_73 = lean_ctor_get(x_57, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_57); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +lean_object* x_74; lean_object* x_75; lean_object* x_76; +x_74 = lean_ctor_get(x_59, 0); +x_75 = lean_ctor_get(x_59, 1); +lean_inc(x_75); +lean_inc(x_74); +lean_dec(x_59); +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_75; +uint8_t x_77; lean_free_object(x_12); lean_dec(x_15); lean_dec(x_14); @@ -20961,170 +20876,141 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_3); lean_dec(x_1); -x_75 = !lean_is_exclusive(x_16); -if (x_75 == 0) +x_77 = !lean_is_exclusive(x_18); +if (x_77 == 0) { -return x_16; +return x_18; } else { -lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_76 = lean_ctor_get(x_16, 0); -x_77 = lean_ctor_get(x_16, 1); -lean_inc(x_77); -lean_inc(x_76); -lean_dec(x_16); -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_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_18, 0); +x_79 = lean_ctor_get(x_18, 1); +lean_inc(x_79); +lean_inc(x_78); +lean_dec(x_18); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +return x_80; } } } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_12, 0); -x_80 = lean_ctor_get(x_12, 1); -lean_inc(x_80); -lean_inc(x_79); +lean_object* x_81; lean_object* x_82; lean_object* x_83; uint8_t x_84; lean_object* x_85; +x_81 = lean_ctor_get(x_12, 0); +x_82 = lean_ctor_get(x_12, 1); +lean_inc(x_82); +lean_inc(x_81); lean_dec(x_12); +lean_inc(x_1); +x_83 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); +lean_closure_set(x_83, 0, x_1); +x_84 = 2; 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_1); -x_81 = l_Lean_Elab_Term_elabType(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_81) == 0) +x_85 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_83, x_84, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_85) == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_82 = lean_ctor_get(x_81, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -lean_dec(x_81); -x_84 = l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__3; -x_85 = l_Lean_Elab_Term_registerCustomErrorIfMVar(x_82, x_1, x_84, x_5, x_6, x_7, x_8, x_9, x_10, x_83); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_86 = lean_ctor_get(x_85, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_85, 1); +lean_inc(x_87); +lean_dec(x_85); +x_88 = l_Lean_Elab_Term_elabLetDeclAux___lambda__1___closed__3; +x_89 = l_Lean_Elab_Term_registerCustomErrorIfMVar(x_86, x_1, x_88, x_5, x_6, x_7, x_8, x_9, x_10, x_87); if (x_2 == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; -x_86 = lean_ctor_get(x_85, 1); +lean_object* x_90; lean_object* x_91; lean_object* x_92; uint8_t x_93; lean_object* x_94; +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +lean_dec(x_89); lean_inc(x_86); -lean_dec(x_85); -lean_inc(x_82); -x_87 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_87, 0, x_82); -x_88 = lean_box(0); -x_89 = 1; +x_91 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_91, 0, x_86); +x_92 = lean_box(0); +x_93 = 1; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_90 = l_Lean_Elab_Term_elabTermEnsuringType(x_3, x_87, x_89, x_89, x_88, x_5, x_6, x_7, x_8, x_9, x_10, x_86); -if (lean_obj_tag(x_90) == 0) +x_94 = l_Lean_Elab_Term_elabTermEnsuringType(x_3, x_91, x_93, x_93, x_92, x_5, x_6, x_7, x_8, x_9, x_10, x_90); +if (lean_obj_tag(x_94) == 0) { -lean_object* x_91; lean_object* x_92; uint8_t x_93; uint8_t x_94; lean_object* x_95; -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 = 0; -x_94 = 1; -lean_inc(x_80); -x_95 = l_Lean_Meta_mkForallFVars(x_80, x_82, x_93, x_89, x_94, x_7, x_8, x_9, x_10, x_92); -if (lean_obj_tag(x_95) == 0) -{ -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_95, 0); +lean_object* x_95; lean_object* x_96; uint8_t x_97; uint8_t x_98; lean_object* x_99; +x_95 = lean_ctor_get(x_94, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_94, 1); lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); -lean_inc(x_97); -lean_dec(x_95); -x_98 = l_Lean_Meta_mkLambdaFVars(x_80, x_91, x_93, x_93, x_94, x_7, x_8, x_9, x_10, x_97); +lean_dec(x_94); +x_97 = 0; +x_98 = 1; +lean_inc(x_82); +x_99 = l_Lean_Meta_mkForallFVars(x_82, x_86, x_97, x_93, x_98, x_7, x_8, x_9, x_10, x_96); +if (lean_obj_tag(x_99) == 0) +{ +lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); +lean_inc(x_101); +lean_dec(x_99); +x_102 = l_Lean_Meta_mkLambdaFVars(x_82, x_95, x_97, x_97, x_98, x_7, x_8, x_9, x_10, x_101); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -if (lean_obj_tag(x_98) == 0) -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_99 = lean_ctor_get(x_98, 0); -lean_inc(x_99); -x_100 = lean_ctor_get(x_98, 1); -lean_inc(x_100); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_101 = x_98; -} else { - lean_dec_ref(x_98); - x_101 = lean_box(0); -} -x_102 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_102, 0, x_99); -lean_ctor_set(x_102, 1, x_79); -x_103 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_103, 0, x_96); -lean_ctor_set(x_103, 1, x_102); -if (lean_is_scalar(x_101)) { - x_104 = lean_alloc_ctor(0, 2, 0); -} else { - x_104 = x_101; -} -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_100); -return x_104; -} -else +if (lean_obj_tag(x_102) == 0) { -lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; -lean_dec(x_96); -lean_dec(x_79); -x_105 = lean_ctor_get(x_98, 0); -lean_inc(x_105); -x_106 = lean_ctor_get(x_98, 1); -lean_inc(x_106); -if (lean_is_exclusive(x_98)) { - lean_ctor_release(x_98, 0); - lean_ctor_release(x_98, 1); - x_107 = x_98; +lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); +lean_inc(x_104); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_105 = x_102; } else { - lean_dec_ref(x_98); - x_107 = lean_box(0); + lean_dec_ref(x_102); + x_105 = lean_box(0); } -if (lean_is_scalar(x_107)) { - x_108 = lean_alloc_ctor(1, 2, 0); +x_106 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_106, 0, x_103); +lean_ctor_set(x_106, 1, x_81); +x_107 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_107, 0, x_100); +lean_ctor_set(x_107, 1, x_106); +if (lean_is_scalar(x_105)) { + x_108 = lean_alloc_ctor(0, 2, 0); } else { - x_108 = x_107; + x_108 = x_105; } -lean_ctor_set(x_108, 0, x_105); -lean_ctor_set(x_108, 1, x_106); +lean_ctor_set(x_108, 0, x_107); +lean_ctor_set(x_108, 1, x_104); return x_108; } -} else { lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -lean_dec(x_91); -lean_dec(x_80); -lean_dec(x_79); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -x_109 = lean_ctor_get(x_95, 0); +lean_dec(x_100); +lean_dec(x_81); +x_109 = lean_ctor_get(x_102, 0); lean_inc(x_109); -x_110 = lean_ctor_get(x_95, 1); +x_110 = lean_ctor_get(x_102, 1); lean_inc(x_110); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_111 = x_95; +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_111 = x_102; } else { - lean_dec_ref(x_95); + lean_dec_ref(x_102); x_111 = lean_box(0); } if (lean_is_scalar(x_111)) { @@ -21140,23 +21026,23 @@ return x_112; else { lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +lean_dec(x_95); lean_dec(x_82); -lean_dec(x_80); -lean_dec(x_79); +lean_dec(x_81); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); -x_113 = lean_ctor_get(x_90, 0); +x_113 = lean_ctor_get(x_99, 0); lean_inc(x_113); -x_114 = lean_ctor_get(x_90, 1); +x_114 = lean_ctor_get(x_99, 1); lean_inc(x_114); -if (lean_is_exclusive(x_90)) { - lean_ctor_release(x_90, 0); - lean_ctor_release(x_90, 1); - x_115 = x_90; +if (lean_is_exclusive(x_99)) { + lean_ctor_release(x_99, 0); + lean_ctor_release(x_99, 1); + x_115 = x_99; } else { - lean_dec_ref(x_90); + lean_dec_ref(x_99); x_115 = lean_box(0); } if (lean_is_scalar(x_115)) { @@ -21171,115 +21057,111 @@ return x_116; } else { -lean_object* x_117; uint8_t x_118; uint8_t x_119; uint8_t x_120; lean_object* x_121; -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_3); -x_117 = lean_ctor_get(x_85, 1); -lean_inc(x_117); -lean_dec(x_85); -x_118 = 0; -x_119 = 1; -x_120 = 1; -x_121 = l_Lean_Meta_mkForallFVars(x_80, x_82, x_118, x_119, x_120, x_7, x_8, x_9, x_10, x_117); -if (lean_obj_tag(x_121) == 0) -{ -lean_object* x_122; lean_object* x_123; lean_object* x_124; uint8_t 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 = 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); -lean_inc(x_122); -x_124 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_124, 0, x_122); -x_125 = 0; -x_126 = lean_box(0); -x_127 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_124, x_125, x_126, x_7, x_8, x_9, x_10, x_123); +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +lean_dec(x_86); +lean_dec(x_82); +lean_dec(x_81); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -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; +lean_dec(x_7); +x_117 = lean_ctor_get(x_94, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_94, 1); +lean_inc(x_118); +if (lean_is_exclusive(x_94)) { + lean_ctor_release(x_94, 0); + lean_ctor_release(x_94, 1); + x_119 = x_94; } else { - lean_dec_ref(x_127); - x_130 = lean_box(0); + lean_dec_ref(x_94); + x_119 = lean_box(0); } -x_131 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_131, 0, x_128); -lean_ctor_set(x_131, 1, x_79); -x_132 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_132, 0, x_122); -lean_ctor_set(x_132, 1, x_131); -if (lean_is_scalar(x_130)) { - x_133 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_119)) { + x_120 = lean_alloc_ctor(1, 2, 0); } else { - x_133 = x_130; + x_120 = x_119; +} +lean_ctor_set(x_120, 0, x_117); +lean_ctor_set(x_120, 1, x_118); +return x_120; } -lean_ctor_set(x_133, 0, x_132); -lean_ctor_set(x_133, 1, x_129); -return x_133; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; -lean_dec(x_79); +lean_object* x_121; uint8_t x_122; uint8_t x_123; uint8_t x_124; lean_object* x_125; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_121 = lean_ctor_get(x_89, 1); +lean_inc(x_121); +lean_dec(x_89); +x_122 = 0; +x_123 = 1; +x_124 = 1; +x_125 = l_Lean_Meta_mkForallFVars(x_82, x_86, x_122, x_123, x_124, x_7, x_8, x_9, x_10, x_121); +if (lean_obj_tag(x_125) == 0) +{ +lean_object* x_126; lean_object* x_127; lean_object* x_128; uint8_t x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_126 = lean_ctor_get(x_125, 0); +lean_inc(x_126); +x_127 = lean_ctor_get(x_125, 1); +lean_inc(x_127); +lean_dec(x_125); +lean_inc(x_126); +x_128 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_128, 0, x_126); +x_129 = 0; +x_130 = lean_box(0); +x_131 = l___private_Lean_Meta_Basic_0__Lean_Meta_mkFreshExprMVarImpl(x_128, x_129, x_130, x_7, x_8, x_9, x_10, x_127); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); -x_134 = lean_ctor_get(x_121, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_121, 1); -lean_inc(x_135); -if (lean_is_exclusive(x_121)) { - lean_ctor_release(x_121, 0); - lean_ctor_release(x_121, 1); - x_136 = x_121; +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_ctor_get(x_131, 1); +lean_inc(x_133); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_134 = x_131; } else { - lean_dec_ref(x_121); - x_136 = lean_box(0); -} -if (lean_is_scalar(x_136)) { - x_137 = lean_alloc_ctor(1, 2, 0); + lean_dec_ref(x_131); + x_134 = lean_box(0); +} +x_135 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_135, 0, x_132); +lean_ctor_set(x_135, 1, x_81); +x_136 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_136, 0, x_126); +lean_ctor_set(x_136, 1, x_135); +if (lean_is_scalar(x_134)) { + x_137 = lean_alloc_ctor(0, 2, 0); } else { - x_137 = x_136; + x_137 = x_134; } -lean_ctor_set(x_137, 0, x_134); -lean_ctor_set(x_137, 1, x_135); +lean_ctor_set(x_137, 0, x_136); +lean_ctor_set(x_137, 1, x_133); return x_137; } -} -} else { lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; -lean_dec(x_80); -lean_dec(x_79); +lean_dec(x_81); 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_138 = lean_ctor_get(x_81, 0); +x_138 = lean_ctor_get(x_125, 0); lean_inc(x_138); -x_139 = lean_ctor_get(x_81, 1); +x_139 = lean_ctor_get(x_125, 1); lean_inc(x_139); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_140 = x_81; +if (lean_is_exclusive(x_125)) { + lean_ctor_release(x_125, 0); + lean_ctor_release(x_125, 1); + x_140 = x_125; } else { - lean_dec_ref(x_81); + lean_dec_ref(x_125); x_140 = lean_box(0); } if (lean_is_scalar(x_140)) { @@ -21293,6 +21175,42 @@ return x_141; } } } +else +{ +lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; +lean_dec(x_82); +lean_dec(x_81); +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_142 = lean_ctor_get(x_85, 0); +lean_inc(x_142); +x_143 = lean_ctor_get(x_85, 1); +lean_inc(x_143); +if (lean_is_exclusive(x_85)) { + lean_ctor_release(x_85, 0); + lean_ctor_release(x_85, 1); + x_144 = x_85; +} else { + lean_dec_ref(x_85); + x_144 = lean_box(0); +} +if (lean_is_scalar(x_144)) { + x_145 = lean_alloc_ctor(1, 2, 0); +} else { + x_145 = x_144; +} +lean_ctor_set(x_145, 0, x_142); +lean_ctor_set(x_145, 1, x_143); +return x_145; +} +} +} } LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabLetDeclAux___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: @@ -23238,7 +23156,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(750u); +x_1 = lean_unsigned_to_nat(773u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23250,7 +23168,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(751u); +x_1 = lean_unsigned_to_nat(774u); x_2 = lean_unsigned_to_nat(129u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23278,7 +23196,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(750u); +x_1 = lean_unsigned_to_nat(773u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23290,7 +23208,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDecl_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(750u); +x_1 = lean_unsigned_to_nat(773u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23409,7 +23327,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(776u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23421,7 +23339,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(754u); +x_1 = lean_unsigned_to_nat(777u); x_2 = lean_unsigned_to_nat(130u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23449,7 +23367,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(776u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23461,7 +23379,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetFunDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(753u); +x_1 = lean_unsigned_to_nat(776u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23581,7 +23499,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(756u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23593,7 +23511,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(757u); +x_1 = lean_unsigned_to_nat(780u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23621,7 +23539,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(756u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(39u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23633,7 +23551,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetDelayedDecl_declR _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(756u); +x_1 = lean_unsigned_to_nat(779u); x_2 = lean_unsigned_to_nat(57u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23753,7 +23671,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(759u); +x_1 = lean_unsigned_to_nat(782u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23765,7 +23683,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(760u); +x_1 = lean_unsigned_to_nat(783u); x_2 = lean_unsigned_to_nat(128u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23793,7 +23711,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(759u); +x_1 = lean_unsigned_to_nat(782u); x_2 = lean_unsigned_to_nat(35u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23805,7 +23723,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(759u); +x_1 = lean_unsigned_to_nat(782u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -23851,7 +23769,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23861,7 +23779,7 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -23871,27 +23789,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2; x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5() { _start: { lean_object* x_1; @@ -23899,17 +23817,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7() { _start: { lean_object* x_1; @@ -23917,37 +23835,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8; x_2 = l___private_Lean_Elab_Binders_0__Lean_Elab_Term_expandBinderIdent___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9; x_2 = l_Lean_Elab_Term_declareTacticSyntax___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11() { _start: { lean_object* x_1; @@ -23955,17 +23873,17 @@ x_1 = lean_mk_string_from_bytes("Binders", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13() { _start: { lean_object* x_1; @@ -23973,33 +23891,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14; -x_2 = lean_unsigned_to_nat(10159u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14; +x_2 = lean_unsigned_to_nat(10163u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -24007,6 +23925,7 @@ return x_5; lean_object* initialize_Lean_Elab_Quotation_Precheck(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_Term(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_BindersUtil(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_Elab_SyntheticMVars(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Elab_PreDefinition_WF_TerminationHint(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Elab_Binders(uint8_t builtin, lean_object* w) { @@ -24022,6 +23941,9 @@ lean_dec_ref(res); res = initialize_Lean_Elab_BindersUtil(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_Elab_SyntheticMVars(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Elab_PreDefinition_WF_TerminationHint(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); @@ -24781,37 +24703,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange___cl if (builtin) {res = l___regBuiltin_Lean_Elab_Term_elabLetTmpDecl_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159____closed__15); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10159_(lean_io_mk_world()); +}l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163____closed__15); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_Binders___hyg_10163_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/BuiltinCommand.c b/stage0/stdlib/Lean/Elab/BuiltinCommand.c index 350a8fd6321c..75a47d4ab4ba 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinCommand.c +++ b/stage0/stdlib/Lean/Elab/BuiltinCommand.c @@ -697,7 +697,7 @@ lean_object* l_Lean_Elab_realizeGlobalConstWithInfos(lean_object*, lean_object*, lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabVariable___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_add_alias(lean_object*, lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, lean_object*); +lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinCommand_0__Lean_Elab_Command_replaceBinderAnnotation___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabReduce(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Command_elabEoi___closed__3; @@ -5421,7 +5421,7 @@ return x_29; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabInitQuot___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_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_4 = lean_st_ref_get(x_2, x_3); x_5 = lean_ctor_get(x_4, 0); lean_inc(x_5); @@ -5431,44 +5431,59 @@ lean_dec(x_4); x_7 = lean_ctor_get(x_5, 0); lean_inc(x_7); lean_dec(x_5); -x_8 = lean_box(4); -x_9 = lean_add_decl(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; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; -x_10 = lean_ctor_get(x_9, 0); +x_8 = lean_st_ref_get(x_2, 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 = lean_ctor_get(x_9, 2); +lean_inc(x_11); lean_dec(x_9); -x_11 = lean_st_ref_get(x_2, x_6); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); +x_12 = l_Lean_Elab_Command_instInhabitedScope; +x_13 = l_List_head_x21___rarg(x_12, x_11); lean_dec(x_11); -x_14 = lean_ctor_get(x_12, 2); +x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Elab_Command_instInhabitedScope; -x_16 = l_List_head_x21___rarg(x_15, x_14); +lean_dec(x_13); +x_15 = lean_box(4); +x_16 = l_Lean_Environment_addDecl(x_7, x_14, x_15); lean_dec(x_14); -x_17 = lean_ctor_get(x_16, 1); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); lean_dec(x_16); -x_18 = l_Lean_KernelException_toMessageData(x_10, x_17); -x_19 = l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(x_18, x_1, x_2, x_13); +x_18 = lean_st_ref_get(x_2, x_10); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); +x_21 = lean_ctor_get(x_19, 2); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_List_head_x21___rarg(x_12, x_21); +lean_dec(x_21); +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Lean_KernelException_toMessageData(x_17, x_23); +x_25 = l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(x_24, x_1, x_2, x_20); lean_dec(x_2); -return x_19; +return x_25; } else { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_9, 0); -lean_inc(x_20); -lean_dec(x_9); -x_21 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_20, x_1, x_2, x_6); +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_16, 0); +lean_inc(x_26); +lean_dec(x_16); +x_27 = l_Lean_setEnv___at_Lean_Elab_Command_elabInitQuot___spec__1(x_26, x_1, x_2, x_10); lean_dec(x_2); lean_dec(x_1); -return x_21; +return x_27; } } } @@ -17249,7 +17264,7 @@ lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = 0; +x_17 = 1; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -18057,241 +18072,243 @@ lean_inc(x_4); x_12 = l_Lean_Elab_Term_elabTerm(x_1, x_2, x_11, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(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; uint8_t x_25; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = l_Lean_Elab_Command_elabCheckCore___lambda__3___closed__1; -x_22 = l_Lean_Elab_Term_levelMVarToParam(x_19, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_20); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); +lean_dec(x_19); +x_22 = l_Lean_Elab_Command_elabCheckCore___lambda__3___closed__1; +x_23 = l_Lean_Elab_Term_levelMVarToParam(x_20, x_22, x_4, x_5, x_6, x_7, x_8, x_9, x_21); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = !lean_is_exclusive(x_8); -if (x_25 == 0) +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_8); +if (x_26 == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; -x_26 = lean_ctor_get(x_8, 2); -x_27 = l_Lean_Elab_Command_elabReduce___lambda__1___closed__2; -x_28 = l_Lean_KVMap_setBool(x_26, x_27, x_15); -lean_ctor_set(x_8, 2, x_28); -x_29 = lean_ctor_get(x_6, 0); -lean_inc(x_29); -x_30 = lean_ctor_get(x_6, 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; uint8_t x_36; +x_27 = lean_ctor_get(x_8, 2); +x_28 = l_Lean_Elab_Command_elabReduce___lambda__1___closed__2; +x_29 = l_Lean_KVMap_setBool(x_27, x_28, x_16); +lean_ctor_set(x_8, 2, x_29); +x_30 = lean_ctor_get(x_6, 0); lean_inc(x_30); -x_31 = lean_ctor_get(x_6, 2); +x_31 = lean_ctor_get(x_6, 1); lean_inc(x_31); -x_32 = lean_ctor_get(x_6, 3); +x_32 = lean_ctor_get(x_6, 2); lean_inc(x_32); -x_33 = lean_ctor_get(x_6, 4); +x_33 = lean_ctor_get(x_6, 3); lean_inc(x_33); -x_34 = lean_ctor_get(x_6, 5); +x_34 = lean_ctor_get(x_6, 4); lean_inc(x_34); -x_35 = !lean_is_exclusive(x_29); -if (x_35 == 0) +x_35 = lean_ctor_get(x_6, 5); +lean_inc(x_35); +x_36 = !lean_is_exclusive(x_30); +if (x_36 == 0) { -uint8_t x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; -x_36 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); -x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); -x_38 = 0; -lean_ctor_set_uint8(x_29, 9, x_38); -x_39 = lean_alloc_ctor(0, 6, 2); -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_uint8(x_39, sizeof(void*)*6, x_36); -lean_ctor_set_uint8(x_39, sizeof(void*)*6 + 1, x_37); +uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; +x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); +x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); +x_39 = 0; +lean_ctor_set_uint8(x_30, 9, x_39); +x_40 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_40, 0, x_30); +lean_ctor_set(x_40, 1, x_31); +lean_ctor_set(x_40, 2, x_32); +lean_ctor_set(x_40, 3, x_33); +lean_ctor_set(x_40, 4, x_34); +lean_ctor_set(x_40, 5, x_35); +lean_ctor_set_uint8(x_40, sizeof(void*)*6, x_37); +lean_ctor_set_uint8(x_40, sizeof(void*)*6 + 1, x_38); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_40 = l_Lean_Meta_reduce(x_23, x_11, x_15, x_15, x_39, x_7, x_8, x_9, x_24); -if (lean_obj_tag(x_40) == 0) +x_41 = l_Lean_Meta_reduce(x_24, x_11, x_16, x_16, x_40, x_7, x_8, x_9, x_25); +if (lean_obj_tag(x_41) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; lean_object* x_45; -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; uint8_t x_45; lean_object* x_46; +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_40); -x_43 = l_Lean_MessageData_ofExpr(x_41); -x_44 = 0; -x_45 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_43, x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_42); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +x_44 = l_Lean_MessageData_ofExpr(x_42); +x_45 = 0; +x_46 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_44, x_45, x_4, x_5, x_6, x_7, x_8, x_9, x_43); 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_45; +return x_46; } else { -uint8_t x_46; +uint8_t x_47; lean_dec(x_8); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_46 = !lean_is_exclusive(x_40); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_41); +if (x_47 == 0) { -return x_40; +return x_41; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_40, 0); -x_48 = lean_ctor_get(x_40, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_41, 0); +x_49 = lean_ctor_get(x_41, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_40); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_41); +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_50; uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_50 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); -x_51 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); -x_52 = lean_ctor_get_uint8(x_29, 0); -x_53 = lean_ctor_get_uint8(x_29, 1); -x_54 = lean_ctor_get_uint8(x_29, 2); -x_55 = lean_ctor_get_uint8(x_29, 3); -x_56 = lean_ctor_get_uint8(x_29, 4); -x_57 = lean_ctor_get_uint8(x_29, 5); -x_58 = lean_ctor_get_uint8(x_29, 6); -x_59 = lean_ctor_get_uint8(x_29, 7); -x_60 = lean_ctor_get_uint8(x_29, 8); -x_61 = lean_ctor_get_uint8(x_29, 10); -x_62 = lean_ctor_get_uint8(x_29, 11); -x_63 = lean_ctor_get_uint8(x_29, 12); -lean_dec(x_29); -x_64 = 0; -x_65 = lean_alloc_ctor(0, 0, 13); -lean_ctor_set_uint8(x_65, 0, x_52); -lean_ctor_set_uint8(x_65, 1, x_53); -lean_ctor_set_uint8(x_65, 2, x_54); -lean_ctor_set_uint8(x_65, 3, x_55); -lean_ctor_set_uint8(x_65, 4, x_56); -lean_ctor_set_uint8(x_65, 5, x_57); -lean_ctor_set_uint8(x_65, 6, x_58); -lean_ctor_set_uint8(x_65, 7, x_59); -lean_ctor_set_uint8(x_65, 8, x_60); -lean_ctor_set_uint8(x_65, 9, x_64); -lean_ctor_set_uint8(x_65, 10, x_61); -lean_ctor_set_uint8(x_65, 11, x_62); -lean_ctor_set_uint8(x_65, 12, x_63); -x_66 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_66, 0, x_65); -lean_ctor_set(x_66, 1, x_30); -lean_ctor_set(x_66, 2, x_31); -lean_ctor_set(x_66, 3, x_32); -lean_ctor_set(x_66, 4, x_33); -lean_ctor_set(x_66, 5, x_34); -lean_ctor_set_uint8(x_66, sizeof(void*)*6, x_50); -lean_ctor_set_uint8(x_66, sizeof(void*)*6 + 1, x_51); +uint8_t x_51; uint8_t x_52; uint8_t x_53; uint8_t x_54; uint8_t x_55; uint8_t x_56; uint8_t x_57; uint8_t x_58; uint8_t x_59; uint8_t x_60; uint8_t x_61; uint8_t x_62; uint8_t x_63; uint8_t x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; +x_51 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); +x_52 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); +x_53 = lean_ctor_get_uint8(x_30, 0); +x_54 = lean_ctor_get_uint8(x_30, 1); +x_55 = lean_ctor_get_uint8(x_30, 2); +x_56 = lean_ctor_get_uint8(x_30, 3); +x_57 = lean_ctor_get_uint8(x_30, 4); +x_58 = lean_ctor_get_uint8(x_30, 5); +x_59 = lean_ctor_get_uint8(x_30, 6); +x_60 = lean_ctor_get_uint8(x_30, 7); +x_61 = lean_ctor_get_uint8(x_30, 8); +x_62 = lean_ctor_get_uint8(x_30, 10); +x_63 = lean_ctor_get_uint8(x_30, 11); +x_64 = lean_ctor_get_uint8(x_30, 12); +lean_dec(x_30); +x_65 = 0; +x_66 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_66, 0, x_53); +lean_ctor_set_uint8(x_66, 1, x_54); +lean_ctor_set_uint8(x_66, 2, x_55); +lean_ctor_set_uint8(x_66, 3, x_56); +lean_ctor_set_uint8(x_66, 4, x_57); +lean_ctor_set_uint8(x_66, 5, x_58); +lean_ctor_set_uint8(x_66, 6, x_59); +lean_ctor_set_uint8(x_66, 7, x_60); +lean_ctor_set_uint8(x_66, 8, x_61); +lean_ctor_set_uint8(x_66, 9, x_65); +lean_ctor_set_uint8(x_66, 10, x_62); +lean_ctor_set_uint8(x_66, 11, x_63); +lean_ctor_set_uint8(x_66, 12, x_64); +x_67 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_31); +lean_ctor_set(x_67, 2, x_32); +lean_ctor_set(x_67, 3, x_33); +lean_ctor_set(x_67, 4, x_34); +lean_ctor_set(x_67, 5, x_35); +lean_ctor_set_uint8(x_67, sizeof(void*)*6, x_51); +lean_ctor_set_uint8(x_67, sizeof(void*)*6 + 1, x_52); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_67 = l_Lean_Meta_reduce(x_23, x_11, x_15, x_15, x_66, x_7, x_8, x_9, x_24); -if (lean_obj_tag(x_67) == 0) +x_68 = l_Lean_Meta_reduce(x_24, x_11, x_16, x_16, x_67, x_7, x_8, x_9, x_25); +if (lean_obj_tag(x_68) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_object* x_72; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; +x_69 = lean_ctor_get(x_68, 0); lean_inc(x_69); -lean_dec(x_67); -x_70 = l_Lean_MessageData_ofExpr(x_68); -x_71 = 0; -x_72 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_70, x_71, x_4, x_5, x_6, x_7, x_8, x_9, x_69); +x_70 = lean_ctor_get(x_68, 1); +lean_inc(x_70); +lean_dec(x_68); +x_71 = l_Lean_MessageData_ofExpr(x_69); +x_72 = 0; +x_73 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_71, x_72, x_4, x_5, x_6, x_7, x_8, x_9, x_70); 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_72; +return x_73; } else { -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_dec(x_8); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_73 = lean_ctor_get(x_67, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_67, 1); +x_74 = lean_ctor_get(x_68, 0); lean_inc(x_74); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_75 = x_67; +x_75 = lean_ctor_get(x_68, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_68)) { + lean_ctor_release(x_68, 0); + lean_ctor_release(x_68, 1); + x_76 = x_68; } else { - lean_dec_ref(x_67); - x_75 = lean_box(0); + lean_dec_ref(x_68); + x_76 = lean_box(0); } -if (lean_is_scalar(x_75)) { - x_76 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_76)) { + x_77 = lean_alloc_ctor(1, 2, 0); } else { - x_76 = x_75; + x_77 = x_76; } -lean_ctor_set(x_76, 0, x_73); -lean_ctor_set(x_76, 1, x_74); -return 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_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; uint8_t x_88; uint8_t x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; uint8_t x_99; uint8_t x_100; uint8_t 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; -x_77 = lean_ctor_get(x_8, 0); -x_78 = lean_ctor_get(x_8, 1); -x_79 = lean_ctor_get(x_8, 2); -x_80 = lean_ctor_get(x_8, 3); -x_81 = lean_ctor_get(x_8, 4); -x_82 = lean_ctor_get(x_8, 5); -x_83 = lean_ctor_get(x_8, 6); -x_84 = lean_ctor_get(x_8, 7); -x_85 = lean_ctor_get(x_8, 8); -x_86 = lean_ctor_get(x_8, 9); -x_87 = lean_ctor_get(x_8, 10); -x_88 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -x_89 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t 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; lean_object* x_98; lean_object* x_99; uint8_t x_100; uint8_t 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; uint8_t x_113; lean_object* x_114; uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; +x_78 = lean_ctor_get(x_8, 0); +x_79 = lean_ctor_get(x_8, 1); +x_80 = lean_ctor_get(x_8, 2); +x_81 = lean_ctor_get(x_8, 3); +x_82 = lean_ctor_get(x_8, 4); +x_83 = lean_ctor_get(x_8, 5); +x_84 = lean_ctor_get(x_8, 6); +x_85 = lean_ctor_get(x_8, 7); +x_86 = lean_ctor_get(x_8, 8); +x_87 = lean_ctor_get(x_8, 9); +x_88 = lean_ctor_get(x_8, 10); +x_89 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +x_90 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); +lean_inc(x_88); lean_inc(x_87); lean_inc(x_86); lean_inc(x_85); @@ -18302,142 +18319,141 @@ lean_inc(x_81); lean_inc(x_80); lean_inc(x_79); lean_inc(x_78); -lean_inc(x_77); lean_dec(x_8); -x_90 = l_Lean_Elab_Command_elabReduce___lambda__1___closed__2; -x_91 = l_Lean_KVMap_setBool(x_79, x_90, x_15); -x_92 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_92, 0, x_77); -lean_ctor_set(x_92, 1, x_78); -lean_ctor_set(x_92, 2, x_91); -lean_ctor_set(x_92, 3, x_80); -lean_ctor_set(x_92, 4, x_81); -lean_ctor_set(x_92, 5, x_82); -lean_ctor_set(x_92, 6, x_83); -lean_ctor_set(x_92, 7, x_84); -lean_ctor_set(x_92, 8, x_85); -lean_ctor_set(x_92, 9, x_86); -lean_ctor_set(x_92, 10, x_87); -lean_ctor_set_uint8(x_92, sizeof(void*)*11, x_88); -lean_ctor_set_uint8(x_92, sizeof(void*)*11 + 1, x_89); -x_93 = lean_ctor_get(x_6, 0); -lean_inc(x_93); -x_94 = lean_ctor_get(x_6, 1); +x_91 = l_Lean_Elab_Command_elabReduce___lambda__1___closed__2; +x_92 = l_Lean_KVMap_setBool(x_80, x_91, x_16); +x_93 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_93, 0, x_78); +lean_ctor_set(x_93, 1, x_79); +lean_ctor_set(x_93, 2, x_92); +lean_ctor_set(x_93, 3, x_81); +lean_ctor_set(x_93, 4, x_82); +lean_ctor_set(x_93, 5, x_83); +lean_ctor_set(x_93, 6, x_84); +lean_ctor_set(x_93, 7, x_85); +lean_ctor_set(x_93, 8, x_86); +lean_ctor_set(x_93, 9, x_87); +lean_ctor_set(x_93, 10, x_88); +lean_ctor_set_uint8(x_93, sizeof(void*)*11, x_89); +lean_ctor_set_uint8(x_93, sizeof(void*)*11 + 1, x_90); +x_94 = lean_ctor_get(x_6, 0); lean_inc(x_94); -x_95 = lean_ctor_get(x_6, 2); +x_95 = lean_ctor_get(x_6, 1); lean_inc(x_95); -x_96 = lean_ctor_get(x_6, 3); +x_96 = lean_ctor_get(x_6, 2); lean_inc(x_96); -x_97 = lean_ctor_get(x_6, 4); +x_97 = lean_ctor_get(x_6, 3); lean_inc(x_97); -x_98 = lean_ctor_get(x_6, 5); +x_98 = lean_ctor_get(x_6, 4); lean_inc(x_98); -x_99 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); -x_100 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); -x_101 = lean_ctor_get_uint8(x_93, 0); -x_102 = lean_ctor_get_uint8(x_93, 1); -x_103 = lean_ctor_get_uint8(x_93, 2); -x_104 = lean_ctor_get_uint8(x_93, 3); -x_105 = lean_ctor_get_uint8(x_93, 4); -x_106 = lean_ctor_get_uint8(x_93, 5); -x_107 = lean_ctor_get_uint8(x_93, 6); -x_108 = lean_ctor_get_uint8(x_93, 7); -x_109 = lean_ctor_get_uint8(x_93, 8); -x_110 = lean_ctor_get_uint8(x_93, 10); -x_111 = lean_ctor_get_uint8(x_93, 11); -x_112 = lean_ctor_get_uint8(x_93, 12); -if (lean_is_exclusive(x_93)) { - x_113 = x_93; +x_99 = lean_ctor_get(x_6, 5); +lean_inc(x_99); +x_100 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); +x_101 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); +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); +x_113 = lean_ctor_get_uint8(x_94, 12); +if (lean_is_exclusive(x_94)) { + x_114 = x_94; } else { - lean_dec_ref(x_93); - x_113 = lean_box(0); + lean_dec_ref(x_94); + x_114 = lean_box(0); } -x_114 = 0; -if (lean_is_scalar(x_113)) { - x_115 = lean_alloc_ctor(0, 0, 13); +x_115 = 0; +if (lean_is_scalar(x_114)) { + x_116 = lean_alloc_ctor(0, 0, 13); } else { - x_115 = x_113; -} -lean_ctor_set_uint8(x_115, 0, x_101); -lean_ctor_set_uint8(x_115, 1, x_102); -lean_ctor_set_uint8(x_115, 2, x_103); -lean_ctor_set_uint8(x_115, 3, x_104); -lean_ctor_set_uint8(x_115, 4, x_105); -lean_ctor_set_uint8(x_115, 5, x_106); -lean_ctor_set_uint8(x_115, 6, x_107); -lean_ctor_set_uint8(x_115, 7, x_108); -lean_ctor_set_uint8(x_115, 8, x_109); -lean_ctor_set_uint8(x_115, 9, x_114); -lean_ctor_set_uint8(x_115, 10, x_110); -lean_ctor_set_uint8(x_115, 11, x_111); -lean_ctor_set_uint8(x_115, 12, x_112); -x_116 = lean_alloc_ctor(0, 6, 2); -lean_ctor_set(x_116, 0, x_115); -lean_ctor_set(x_116, 1, x_94); -lean_ctor_set(x_116, 2, x_95); -lean_ctor_set(x_116, 3, x_96); -lean_ctor_set(x_116, 4, x_97); -lean_ctor_set(x_116, 5, x_98); -lean_ctor_set_uint8(x_116, sizeof(void*)*6, x_99); -lean_ctor_set_uint8(x_116, sizeof(void*)*6 + 1, x_100); + x_116 = x_114; +} +lean_ctor_set_uint8(x_116, 0, x_102); +lean_ctor_set_uint8(x_116, 1, x_103); +lean_ctor_set_uint8(x_116, 2, x_104); +lean_ctor_set_uint8(x_116, 3, x_105); +lean_ctor_set_uint8(x_116, 4, x_106); +lean_ctor_set_uint8(x_116, 5, x_107); +lean_ctor_set_uint8(x_116, 6, x_108); +lean_ctor_set_uint8(x_116, 7, x_109); +lean_ctor_set_uint8(x_116, 8, x_110); +lean_ctor_set_uint8(x_116, 9, x_115); +lean_ctor_set_uint8(x_116, 10, x_111); +lean_ctor_set_uint8(x_116, 11, x_112); +lean_ctor_set_uint8(x_116, 12, x_113); +x_117 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_117, 0, x_116); +lean_ctor_set(x_117, 1, x_95); +lean_ctor_set(x_117, 2, x_96); +lean_ctor_set(x_117, 3, x_97); +lean_ctor_set(x_117, 4, x_98); +lean_ctor_set(x_117, 5, x_99); +lean_ctor_set_uint8(x_117, sizeof(void*)*6, x_100); +lean_ctor_set_uint8(x_117, sizeof(void*)*6 + 1, x_101); lean_inc(x_9); -lean_inc(x_92); +lean_inc(x_93); lean_inc(x_7); -x_117 = l_Lean_Meta_reduce(x_23, x_11, x_15, x_15, x_116, x_7, x_92, x_9, x_24); -if (lean_obj_tag(x_117) == 0) +x_118 = l_Lean_Meta_reduce(x_24, x_11, x_16, x_16, x_117, x_7, x_93, x_9, x_25); +if (lean_obj_tag(x_118) == 0) { -lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t 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_object* x_119; lean_object* x_120; lean_object* x_121; uint8_t x_122; lean_object* x_123; +x_119 = lean_ctor_get(x_118, 0); lean_inc(x_119); -lean_dec(x_117); -x_120 = l_Lean_MessageData_ofExpr(x_118); -x_121 = 0; -x_122 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_120, x_121, x_4, x_5, x_6, x_7, x_92, x_9, x_119); +x_120 = lean_ctor_get(x_118, 1); +lean_inc(x_120); +lean_dec(x_118); +x_121 = l_Lean_MessageData_ofExpr(x_119); +x_122 = 0; +x_123 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_3, x_121, x_122, x_4, x_5, x_6, x_7, x_93, x_9, x_120); lean_dec(x_9); -lean_dec(x_92); +lean_dec(x_93); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_122; +return x_123; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -lean_dec(x_92); +lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; +lean_dec(x_93); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_123 = lean_ctor_get(x_117, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_117, 1); +x_124 = lean_ctor_get(x_118, 0); lean_inc(x_124); -if (lean_is_exclusive(x_117)) { - lean_ctor_release(x_117, 0); - lean_ctor_release(x_117, 1); - x_125 = x_117; +x_125 = lean_ctor_get(x_118, 1); +lean_inc(x_125); +if (lean_is_exclusive(x_118)) { + lean_ctor_release(x_118, 0); + lean_ctor_release(x_118, 1); + x_126 = x_118; } else { - lean_dec_ref(x_117); - x_125 = lean_box(0); + lean_dec_ref(x_118); + x_126 = lean_box(0); } -if (lean_is_scalar(x_125)) { - x_126 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_126)) { + x_127 = lean_alloc_ctor(1, 2, 0); } else { - x_126 = x_125; + x_127 = x_126; } -lean_ctor_set(x_126, 0, x_123); -lean_ctor_set(x_126, 1, x_124); -return x_126; +lean_ctor_set(x_127, 0, x_124); +lean_ctor_set(x_127, 1, x_125); +return x_127; } } } else { -uint8_t x_127; +uint8_t x_128; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -18445,52 +18461,52 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_127 = !lean_is_exclusive(x_16); -if (x_127 == 0) +x_128 = !lean_is_exclusive(x_17); +if (x_128 == 0) { -return x_16; +return x_17; } else { -lean_object* x_128; lean_object* x_129; lean_object* x_130; -x_128 = lean_ctor_get(x_16, 0); -x_129 = lean_ctor_get(x_16, 1); +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_17, 0); +x_130 = lean_ctor_get(x_17, 1); +lean_inc(x_130); lean_inc(x_129); -lean_inc(x_128); -lean_dec(x_16); -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; +lean_dec(x_17); +x_131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +return x_131; } } } else { -uint8_t x_131; +uint8_t x_132; 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_131 = !lean_is_exclusive(x_12); -if (x_131 == 0) +x_132 = !lean_is_exclusive(x_12); +if (x_132 == 0) { return x_12; } else { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -x_132 = lean_ctor_get(x_12, 0); -x_133 = lean_ctor_get(x_12, 1); +lean_object* x_133; lean_object* x_134; lean_object* x_135; +x_133 = lean_ctor_get(x_12, 0); +x_134 = lean_ctor_get(x_12, 1); +lean_inc(x_134); lean_inc(x_133); -lean_inc(x_132); lean_dec(x_12); -x_134 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_134, 0, x_132); -lean_ctor_set(x_134, 1, x_133); -return x_134; +x_135 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_135, 0, x_133); +lean_ctor_set(x_135, 1, x_134); +return x_135; } } } @@ -22281,148 +22297,149 @@ lean_inc(x_3); x_172 = l_Lean_Elab_Term_elabTerm(x_2, x_3, x_171, x_171, x_4, x_5, x_6, x_7, x_8, x_9, x_13); if (lean_obj_tag(x_172) == 0) { -lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; +lean_object* x_173; lean_object* x_174; uint8_t x_175; uint8_t x_176; lean_object* x_177; x_173 = lean_ctor_get(x_172, 0); lean_inc(x_173); x_174 = lean_ctor_get(x_172, 1); lean_inc(x_174); lean_dec(x_172); -x_175 = 0; +x_175 = 1; +x_176 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_176 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_175, x_175, x_4, x_5, x_6, x_7, x_8, x_9, x_174); -if (lean_obj_tag(x_176) == 0) +x_177 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_175, x_176, x_4, x_5, x_6, x_7, x_8, x_9, x_174); +if (lean_obj_tag(x_177) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 1); -lean_inc(x_177); -lean_dec(x_176); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +lean_dec(x_177); lean_inc(x_173); -x_178 = l_Lean_Meta_getMVars(x_173, x_6, x_7, x_8, x_9, x_177); -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -x_180 = lean_ctor_get(x_178, 1); +x_179 = l_Lean_Meta_getMVars(x_173, x_6, x_7, x_8, x_9, x_178); +x_180 = lean_ctor_get(x_179, 0); lean_inc(x_180); -lean_dec(x_178); +x_181 = lean_ctor_get(x_179, 1); +lean_inc(x_181); +lean_dec(x_179); 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_181 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_179, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_180); -lean_dec(x_179); -if (lean_obj_tag(x_181) == 0) +x_182 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_180, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_181); +lean_dec(x_180); +if (lean_obj_tag(x_182) == 0) { -lean_object* x_182; uint8_t x_183; -x_182 = lean_ctor_get(x_181, 0); -lean_inc(x_182); -x_183 = lean_unbox(x_182); -lean_dec(x_182); -if (x_183 == 0) +lean_object* x_183; uint8_t x_184; +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +x_184 = lean_unbox(x_183); +lean_dec(x_183); +if (x_184 == 0) { -lean_object* x_184; lean_object* x_185; lean_object* x_186; -x_184 = lean_ctor_get(x_181, 1); -lean_inc(x_184); -lean_dec(x_181); -x_185 = lean_box(0); +lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_185 = lean_ctor_get(x_182, 1); +lean_inc(x_185); +lean_dec(x_182); +x_186 = lean_box(0); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_186 = l_Lean_Elab_Command_elabEvalUnsafe___lambda__2(x_173, x_185, x_4, x_5, x_6, x_7, x_8, x_9, x_184); -if (lean_obj_tag(x_186) == 0) +x_187 = l_Lean_Elab_Command_elabEvalUnsafe___lambda__2(x_173, x_186, x_4, x_5, x_6, x_7, x_8, x_9, x_185); +if (lean_obj_tag(x_187) == 0) { -lean_object* x_187; lean_object* x_188; -x_187 = lean_ctor_get(x_186, 0); -lean_inc(x_187); -x_188 = lean_ctor_get(x_186, 1); +lean_object* x_188; lean_object* x_189; +x_188 = lean_ctor_get(x_187, 0); lean_inc(x_188); -lean_dec(x_186); -x_31 = x_187; -x_32 = x_188; +x_189 = lean_ctor_get(x_187, 1); +lean_inc(x_189); +lean_dec(x_187); +x_31 = x_188; +x_32 = x_189; goto block_170; } else { -lean_object* x_189; lean_object* x_190; +lean_object* x_190; lean_object* x_191; lean_dec(x_1); -x_189 = lean_ctor_get(x_186, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_186, 1); +x_190 = lean_ctor_get(x_187, 0); lean_inc(x_190); -lean_dec(x_186); -x_23 = x_189; -x_24 = x_190; +x_191 = lean_ctor_get(x_187, 1); +lean_inc(x_191); +lean_dec(x_187); +x_23 = x_190; +x_24 = x_191; goto block_30; } } else { -lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_dec(x_173); lean_dec(x_1); -x_191 = lean_ctor_get(x_181, 1); -lean_inc(x_191); -lean_dec(x_181); -x_192 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_191); -x_193 = lean_ctor_get(x_192, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_192, 1); +x_192 = lean_ctor_get(x_182, 1); +lean_inc(x_192); +lean_dec(x_182); +x_193 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_192); +x_194 = lean_ctor_get(x_193, 0); lean_inc(x_194); -lean_dec(x_192); -x_23 = x_193; -x_24 = x_194; +x_195 = lean_ctor_get(x_193, 1); +lean_inc(x_195); +lean_dec(x_193); +x_23 = x_194; +x_24 = x_195; goto block_30; } } else { -lean_object* x_195; lean_object* x_196; +lean_object* x_196; lean_object* x_197; lean_dec(x_173); lean_dec(x_1); -x_195 = lean_ctor_get(x_181, 0); -lean_inc(x_195); -x_196 = lean_ctor_get(x_181, 1); +x_196 = lean_ctor_get(x_182, 0); lean_inc(x_196); -lean_dec(x_181); -x_23 = x_195; -x_24 = x_196; +x_197 = lean_ctor_get(x_182, 1); +lean_inc(x_197); +lean_dec(x_182); +x_23 = x_196; +x_24 = x_197; goto block_30; } } else { -lean_object* x_197; lean_object* x_198; +lean_object* x_198; lean_object* x_199; lean_dec(x_173); lean_dec(x_3); lean_dec(x_1); -x_197 = lean_ctor_get(x_176, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_176, 1); +x_198 = lean_ctor_get(x_177, 0); lean_inc(x_198); -lean_dec(x_176); -x_23 = x_197; -x_24 = x_198; +x_199 = lean_ctor_get(x_177, 1); +lean_inc(x_199); +lean_dec(x_177); +x_23 = x_198; +x_24 = x_199; goto block_30; } } else { -lean_object* x_199; lean_object* x_200; +lean_object* x_200; lean_object* x_201; lean_dec(x_3); lean_dec(x_1); -x_199 = lean_ctor_get(x_172, 0); -lean_inc(x_199); -x_200 = lean_ctor_get(x_172, 1); +x_200 = lean_ctor_get(x_172, 0); lean_inc(x_200); +x_201 = lean_ctor_get(x_172, 1); +lean_inc(x_201); lean_dec(x_172); -x_23 = x_199; -x_24 = x_200; +x_23 = x_200; +x_24 = x_201; goto block_30; } block_22: @@ -23025,148 +23042,149 @@ lean_inc(x_4); x_106 = l_Lean_Elab_Term_elabTerm(x_3, x_4, x_105, x_105, x_5, x_6, x_7, x_8, x_9, x_10, x_14); if (lean_obj_tag(x_106) == 0) { -lean_object* x_107; lean_object* x_108; uint8_t x_109; lean_object* x_110; +lean_object* x_107; lean_object* x_108; uint8_t x_109; uint8_t x_110; lean_object* x_111; 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 = 0; +x_109 = 1; +x_110 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_110 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_109, x_109, x_5, x_6, x_7, x_8, x_9, x_10, x_108); -if (lean_obj_tag(x_110) == 0) +x_111 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_109, x_110, x_5, x_6, x_7, x_8, x_9, x_10, x_108); +if (lean_obj_tag(x_111) == 0) { -lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_111 = lean_ctor_get(x_110, 1); -lean_inc(x_111); -lean_dec(x_110); +lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_112 = lean_ctor_get(x_111, 1); +lean_inc(x_112); +lean_dec(x_111); lean_inc(x_107); -x_112 = l_Lean_Meta_getMVars(x_107, x_7, x_8, x_9, x_10, x_111); -x_113 = lean_ctor_get(x_112, 0); -lean_inc(x_113); -x_114 = lean_ctor_get(x_112, 1); +x_113 = l_Lean_Meta_getMVars(x_107, x_7, x_8, x_9, x_10, x_112); +x_114 = lean_ctor_get(x_113, 0); lean_inc(x_114); -lean_dec(x_112); +x_115 = lean_ctor_get(x_113, 1); +lean_inc(x_115); +lean_dec(x_113); 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_115 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_113, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_114); -lean_dec(x_113); -if (lean_obj_tag(x_115) == 0) +x_116 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_114, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_115); +lean_dec(x_114); +if (lean_obj_tag(x_116) == 0) { -lean_object* x_116; uint8_t x_117; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -x_117 = lean_unbox(x_116); -lean_dec(x_116); -if (x_117 == 0) +lean_object* x_117; uint8_t x_118; +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_unbox(x_117); +lean_dec(x_117); +if (x_118 == 0) { -lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_118 = lean_ctor_get(x_115, 1); -lean_inc(x_118); -lean_dec(x_115); -x_119 = lean_box(0); +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_116, 1); +lean_inc(x_119); +lean_dec(x_116); +x_120 = lean_box(0); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_120 = l_Lean_Elab_Command_elabEvalUnsafe___lambda__2(x_107, x_119, x_5, x_6, x_7, x_8, x_9, x_10, x_118); -if (lean_obj_tag(x_120) == 0) +x_121 = l_Lean_Elab_Command_elabEvalUnsafe___lambda__2(x_107, x_120, x_5, x_6, x_7, x_8, x_9, x_10, x_119); +if (lean_obj_tag(x_121) == 0) { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); +lean_object* x_122; lean_object* x_123; +x_122 = lean_ctor_get(x_121, 0); lean_inc(x_122); -lean_dec(x_120); -x_32 = x_121; -x_33 = x_122; +x_123 = lean_ctor_get(x_121, 1); +lean_inc(x_123); +lean_dec(x_121); +x_32 = x_122; +x_33 = x_123; goto block_104; } else { -lean_object* x_123; lean_object* x_124; +lean_object* x_124; lean_object* x_125; lean_dec(x_1); -x_123 = lean_ctor_get(x_120, 0); -lean_inc(x_123); -x_124 = lean_ctor_get(x_120, 1); +x_124 = lean_ctor_get(x_121, 0); lean_inc(x_124); -lean_dec(x_120); -x_24 = x_123; -x_25 = x_124; +x_125 = lean_ctor_get(x_121, 1); +lean_inc(x_125); +lean_dec(x_121); +x_24 = x_124; +x_25 = x_125; goto block_31; } } else { -lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; +lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_dec(x_107); lean_dec(x_1); -x_125 = lean_ctor_get(x_115, 1); -lean_inc(x_125); -lean_dec(x_115); -x_126 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_125); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_ctor_get(x_126, 1); +x_126 = lean_ctor_get(x_116, 1); +lean_inc(x_126); +lean_dec(x_116); +x_127 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_126); +x_128 = lean_ctor_get(x_127, 0); lean_inc(x_128); -lean_dec(x_126); -x_24 = x_127; -x_25 = x_128; +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_24 = x_128; +x_25 = x_129; goto block_31; } } else { -lean_object* x_129; lean_object* x_130; +lean_object* x_130; lean_object* x_131; lean_dec(x_107); lean_dec(x_1); -x_129 = lean_ctor_get(x_115, 0); -lean_inc(x_129); -x_130 = lean_ctor_get(x_115, 1); +x_130 = lean_ctor_get(x_116, 0); lean_inc(x_130); -lean_dec(x_115); -x_24 = x_129; -x_25 = x_130; +x_131 = lean_ctor_get(x_116, 1); +lean_inc(x_131); +lean_dec(x_116); +x_24 = x_130; +x_25 = x_131; goto block_31; } } else { -lean_object* x_131; lean_object* x_132; +lean_object* x_132; lean_object* x_133; lean_dec(x_107); lean_dec(x_4); lean_dec(x_1); -x_131 = lean_ctor_get(x_110, 0); -lean_inc(x_131); -x_132 = lean_ctor_get(x_110, 1); +x_132 = lean_ctor_get(x_111, 0); lean_inc(x_132); -lean_dec(x_110); -x_24 = x_131; -x_25 = x_132; +x_133 = lean_ctor_get(x_111, 1); +lean_inc(x_133); +lean_dec(x_111); +x_24 = x_132; +x_25 = x_133; goto block_31; } } else { -lean_object* x_133; lean_object* x_134; +lean_object* x_134; lean_object* x_135; lean_dec(x_4); lean_dec(x_1); -x_133 = lean_ctor_get(x_106, 0); -lean_inc(x_133); -x_134 = lean_ctor_get(x_106, 1); +x_134 = lean_ctor_get(x_106, 0); lean_inc(x_134); +x_135 = lean_ctor_get(x_106, 1); +lean_inc(x_135); lean_dec(x_106); -x_24 = x_133; -x_25 = x_134; +x_24 = x_134; +x_25 = x_135; goto block_31; } block_23: @@ -25692,109 +25710,110 @@ lean_inc(x_2); x_11 = l_Lean_Elab_Term_elabTerm(x_1, x_2, x_10, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_11) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = 0; +x_14 = 1; +x_15 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_15 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_14, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_14, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +if (lean_obj_tag(x_16) == 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, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_16); -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_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_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_20 = l_Lean_Meta_synthInstance(x_18, x_2, x_5, x_6, x_7, x_8, x_19); -if (lean_obj_tag(x_20) == 0) +x_21 = l_Lean_Meta_synthInstance(x_19, x_2, x_5, x_6, x_7, x_8, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t 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_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; uint8_t x_27; +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = l_Lean_MessageData_ofExpr(x_21); -x_24 = 0; -x_25 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_23, x_24, x_3, x_4, x_5, x_6, x_7, x_8, 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 = 0; +x_26 = l_Lean_log___at_Lean_Elab_Term_exceptionToSorry___spec__2(x_24, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_23); 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_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 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_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_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_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 { -uint8_t x_32; +uint8_t x_33; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_32 = !lean_is_exclusive(x_20); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_21); +if (x_33 == 0) { -return x_20; +return x_21; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_20, 0); -x_34 = lean_ctor_get(x_20, 1); +lean_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_inc(x_33); -lean_dec(x_20); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_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; } } } else { -uint8_t x_36; +uint8_t x_37; lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); @@ -25803,29 +25822,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_36 = !lean_is_exclusive(x_15); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_16); +if (x_37 == 0) { -return x_15; +return x_16; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_15, 0); -x_38 = lean_ctor_get(x_15, 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_15); -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_dec(x_16); +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 { -uint8_t x_40; +uint8_t x_41; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25833,23 +25852,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_40 = !lean_is_exclusive(x_11); -if (x_40 == 0) +x_41 = !lean_is_exclusive(x_11); +if (x_41 == 0) { return x_11; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_11, 0); -x_42 = lean_ctor_get(x_11, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_11, 0); +x_43 = lean_ctor_get(x_11, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); lean_dec(x_11); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/BuiltinNotation.c b/stage0/stdlib/Lean/Elab/BuiltinNotation.c index 47afc0c4e0e8..340ae8cd7b71 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinNotation.c +++ b/stage0/stdlib/Lean/Elab/BuiltinNotation.c @@ -84,6 +84,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabNoindex___closed__2; lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_expandHave___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeAscription___closed__2; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___closed__44; static lean_object* l___regBuiltin_Lean_Elab_Term_elabTrailingParserMacro_declRange___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry(lean_object*); @@ -424,7 +425,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabCoeSortNotation_declRange_ static lean_object* l___regBuiltin_Lean_Elab_Term_elabSorry_declRange___closed__5; static lean_object* l_Lean_Elab_Term_elabRunElab_unsafe__1___closed__7; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Elab_Term_elabCDotFunctionAlias_x3f___spec__21___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkPairs_loop___closed__7; static lean_object* l_Lean_Elab_Term_expandDbgTrace___closed__16; static lean_object* l_Lean_Elab_Term_expandHave___closed__5; @@ -658,7 +658,7 @@ static lean_object* l_Lean_Elab_Term_expandHave___lambda__2___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_expandShow_declRange___closed__6; static lean_object* l_Lean_Elab_Term_expandShow___closed__10; uint8_t l_Lean_LocalDecl_isAuxDecl(lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabShow_declRange___closed__6; static lean_object* l_Lean_Elab_Term_elabPanic___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at___private_Lean_Elab_BuiltinNotation_0__Lean_Elab_Term_elabParserMacroAux___spec__4(lean_object*, size_t, size_t); @@ -2327,7 +2327,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Term_elabAnonymo lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; @@ -4899,14 +4899,14 @@ lean_inc(x_2); x_21 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabShow___lambda__2), 9, 2); lean_closure_set(x_21, 0, x_14); lean_closure_set(x_21, 1, x_2); -x_22 = 1; +x_22 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_23 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_21, x_22, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_23 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_21, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_23) == 0) { lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; @@ -4927,7 +4927,7 @@ lean_inc(x_3); x_29 = l_Lean_Elab_Term_exprToSyntax(x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_25); if (lean_obj_tag(x_29) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); x_31 = lean_ctor_get(x_29, 1); @@ -4981,12 +4981,13 @@ lean_ctor_set(x_53, 0, x_33); lean_ctor_set(x_53, 1, x_52); x_54 = l_Lean_Elab_Term_elabShow___closed__4; x_55 = l_Lean_Syntax_node4(x_33, x_54, x_37, x_51, x_53, x_28); -x_56 = l_Lean_Elab_Term_elabTerm(x_55, x_2, x_22, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_35); -return x_56; +x_56 = 1; +x_57 = l_Lean_Elab_Term_elabTerm(x_55, x_2, x_56, x_56, x_3, x_4, x_5, x_6, x_7, x_8, x_35); +return x_57; } else { -uint8_t x_57; +uint8_t x_58; lean_dec(x_28); lean_dec(x_20); lean_dec(x_8); @@ -4996,29 +4997,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_57 = !lean_is_exclusive(x_29); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_29); +if (x_58 == 0) { return x_29; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_29, 0); -x_59 = lean_ctor_get(x_29, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_29, 0); +x_60 = lean_ctor_get(x_29, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); lean_dec(x_29); -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; +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 { -uint8_t x_61; +uint8_t x_62; lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -5028,23 +5029,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_61 = !lean_is_exclusive(x_23); -if (x_61 == 0) +x_62 = !lean_is_exclusive(x_23); +if (x_62 == 0) { return x_23; } else { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -x_62 = lean_ctor_get(x_23, 0); -x_63 = lean_ctor_get(x_23, 1); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_23, 0); +x_64 = lean_ctor_get(x_23, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); lean_dec(x_23); -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; +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; } } } @@ -19474,14 +19475,14 @@ lean_closure_set(x_25, 0, x_14); lean_closure_set(x_25, 1, x_21); lean_closure_set(x_25, 2, x_23); lean_closure_set(x_25, 3, x_24); -x_26 = 0; +x_26 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_27 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_25, x_26, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_27 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_25, x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_27) == 0) { lean_object* x_28; lean_object* x_29; lean_object* x_30; @@ -19533,17 +19534,17 @@ x_36 = l_Lean_Syntax_getArg(x_16, x_35); lean_dec(x_16); x_37 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); lean_closure_set(x_37, 0, x_36); -x_38 = 1; +x_38 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_39 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_37, x_38, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_39 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_37, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_39) == 0) { -lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; x_40 = lean_ctor_get(x_39, 0); lean_inc(x_40); x_41 = lean_ctor_get(x_39, 1); @@ -19551,6 +19552,7 @@ lean_inc(x_41); lean_dec(x_39); x_42 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_42, 0, x_40); +x_43 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -19558,22 +19560,22 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_42); -x_43 = l_Lean_Elab_Term_elabTerm(x_14, x_42, x_38, x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_41); -if (lean_obj_tag(x_43) == 0) +x_44 = l_Lean_Elab_Term_elabTerm(x_14, x_42, x_43, x_43, x_3, x_4, x_5, x_6, x_7, x_8, x_41); +if (lean_obj_tag(x_44) == 0) { -lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); -x_46 = lean_box(0); -x_47 = l_Lean_Elab_Term_ensureHasType(x_42, x_44, x_46, x_46, x_3, x_4, x_5, x_6, x_7, x_8, x_45); -return x_47; +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); +x_47 = lean_box(0); +x_48 = l_Lean_Elab_Term_ensureHasType(x_42, x_45, x_47, x_47, x_3, x_4, x_5, x_6, x_7, x_8, x_46); +return x_48; } else { -uint8_t x_48; +uint8_t x_49; lean_dec(x_42); lean_dec(x_8); lean_dec(x_7); @@ -19581,29 +19583,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_48 = !lean_is_exclusive(x_43); -if (x_48 == 0) +x_49 = !lean_is_exclusive(x_44); +if (x_49 == 0) { -return x_43; +return x_44; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_43, 0); -x_50 = lean_ctor_get(x_43, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_44, 0); +x_51 = lean_ctor_get(x_44, 1); +lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); -lean_dec(x_43); -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_dec(x_44); +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_52; +uint8_t x_53; lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); @@ -19611,23 +19613,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_52 = !lean_is_exclusive(x_39); -if (x_52 == 0) +x_53 = !lean_is_exclusive(x_39); +if (x_53 == 0) { return x_39; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_39, 0); -x_54 = lean_ctor_get(x_39, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_39, 0); +x_55 = lean_ctor_get(x_39, 1); +lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); lean_dec(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; +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; } } } @@ -24641,7 +24643,7 @@ else lean_object* x_23; uint8_t x_24; uint8_t x_25; lean_object* x_26; x_23 = l_Lean_Syntax_getArg(x_19, x_16); lean_dec(x_19); -x_24 = 1; +x_24 = 0; x_25 = 0; lean_inc(x_8); lean_inc(x_7); @@ -24652,467 +24654,469 @@ lean_inc(x_3); x_26 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_24, x_25, x_3, x_4, x_5, x_6, x_7, x_8, x_12); if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; uint8_t x_33; lean_object* x_34; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); x_28 = lean_box(0); -x_29 = lean_box(x_24); -x_30 = lean_box(x_24); +x_29 = 1; +x_30 = lean_box(x_29); +x_31 = lean_box(x_29); lean_inc(x_17); -x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_31, 0, x_17); -lean_closure_set(x_31, 1, x_28); -lean_closure_set(x_31, 2, x_29); -lean_closure_set(x_31, 3, x_30); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_32, 0, x_17); +lean_closure_set(x_32, 1, x_28); +lean_closure_set(x_32, 2, x_30); +lean_closure_set(x_32, 3, x_31); +x_33 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_32 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_31, x_25, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_27); -if (lean_obj_tag(x_32) == 0) +x_34 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_32, x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_27); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); -lean_inc(x_34); -lean_dec(x_32); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +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); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_33); -x_35 = lean_infer_type(x_33, x_5, x_6, x_7, x_8, x_34); -if (lean_obj_tag(x_35) == 0) +lean_inc(x_35); +x_37 = lean_infer_type(x_35, x_5, x_6, x_7, x_8, x_36); +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; -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_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_36, x_3, x_4, x_5, x_6, x_7, x_8, x_37); -x_39 = lean_ctor_get(x_38, 0); +lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_ctor_get(x_37, 1); lean_inc(x_39); -x_40 = lean_ctor_get(x_38, 1); -lean_inc(x_40); -lean_dec(x_38); +lean_dec(x_37); +x_40 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_38, x_3, x_4, x_5, x_6, x_7, x_8, x_39); +x_41 = lean_ctor_get(x_40, 0); +lean_inc(x_41); +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); +lean_dec(x_40); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_39); -x_41 = l_Lean_Meta_matchEq_x3f(x_39, x_5, x_6, x_7, x_8, x_40); -if (lean_obj_tag(x_41) == 0) +lean_inc(x_41); +x_43 = l_Lean_Meta_matchEq_x3f(x_41, x_5, x_6, x_7, x_8, x_42); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_42; -x_42 = lean_ctor_get(x_41, 0); -lean_inc(x_42); -if (lean_obj_tag(x_42) == 0) +lean_object* x_44; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +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; +lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; 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_dec(x_23); lean_dec(x_17); lean_dec(x_11); lean_dec(x_1); -x_43 = lean_ctor_get(x_41, 1); -lean_inc(x_43); -lean_dec(x_41); -x_44 = l_Lean_indentExpr(x_33); -x_45 = l_Lean_Elab_Term_elabSubst___closed__3; -x_46 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_44); -x_47 = l_Lean_Elab_Term_elabSubst___closed__5; +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_Lean_indentExpr(x_35); +x_47 = l_Lean_Elab_Term_elabSubst___closed__3; x_48 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_indentExpr(x_39); +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_46); +x_49 = l_Lean_Elab_Term_elabSubst___closed__5; 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_Elab_Term_elabSubst___closed__7; +x_51 = l_Lean_indentExpr(x_41); x_52 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_52, 0, x_50); lean_ctor_set(x_52, 1, x_51); -x_53 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_52, x_3, x_4, x_5, x_6, x_7, x_8, x_43); +x_53 = l_Lean_Elab_Term_elabSubst___closed__7; +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 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_54, x_3, x_4, x_5, x_6, x_7, x_8, x_45); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_53; +return x_55; } else { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_42, 0); -lean_inc(x_54); -lean_dec(x_42); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); +lean_object* x_56; lean_object* x_57; +x_56 = lean_ctor_get(x_44, 0); +lean_inc(x_56); +lean_dec(x_44); +x_57 = lean_ctor_get(x_56, 1); +lean_inc(x_57); if (lean_obj_tag(x_11) == 0) { -lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_dec(x_17); lean_dec(x_1); -x_56 = lean_ctor_get(x_41, 1); -lean_inc(x_56); -lean_dec(x_41); -x_57 = lean_ctor_get(x_54, 0); -lean_inc(x_57); -lean_dec(x_54); -x_58 = lean_ctor_get(x_55, 0); +x_58 = lean_ctor_get(x_43, 1); lean_inc(x_58); -x_59 = lean_ctor_get(x_55, 1); +lean_dec(x_43); +x_59 = lean_ctor_get(x_56, 0); lean_inc(x_59); -lean_dec(x_55); +lean_dec(x_56); +x_60 = lean_ctor_get(x_57, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_57, 1); +lean_inc(x_61); +lean_dec(x_57); 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_60 = l_Lean_Elab_Term_elabTerm(x_23, x_28, x_24, x_24, x_3, x_4, x_5, x_6, x_7, x_8, x_56); -if (lean_obj_tag(x_60) == 0) +x_62 = l_Lean_Elab_Term_elabTerm(x_23, x_28, x_29, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_58); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_60, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_60, 1); -lean_inc(x_62); -lean_dec(x_60); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_ctor_get(x_62, 1); +lean_inc(x_64); +lean_dec(x_62); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_61); -x_63 = lean_infer_type(x_61, x_5, x_6, x_7, x_8, x_62); -if (lean_obj_tag(x_63) == 0) +lean_inc(x_63); +x_65 = lean_infer_type(x_63, x_5, x_6, x_7, x_8, x_64); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_ctor_get(x_63, 1); -lean_inc(x_65); -lean_dec(x_63); -x_66 = lean_box(0); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_58); -lean_inc(x_64); -x_67 = l_Lean_Meta_kabstract(x_64, x_58, x_66, x_5, x_6, x_7, x_8, x_65); -if (lean_obj_tag(x_67) == 0) +lean_inc(x_60); +lean_inc(x_66); +x_69 = l_Lean_Meta_kabstract(x_66, x_60, x_68, x_5, x_6, x_7, x_8, x_67); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -lean_inc(x_61); -lean_inc(x_57); -x_70 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__4___boxed), 14, 2); -lean_closure_set(x_70, 0, x_57); -lean_closure_set(x_70, 1, x_61); -x_71 = l_Lean_Expr_hasLooseBVars(x_68); -if (x_71 == 0) +lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +lean_inc(x_63); +lean_inc(x_59); +x_72 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__4___boxed), 14, 2); +lean_closure_set(x_72, 0, x_59); +lean_closure_set(x_72, 1, x_63); +x_73 = l_Lean_Expr_hasLooseBVars(x_70); +if (x_73 == 0) { -lean_object* x_72; -lean_dec(x_68); -lean_dec(x_61); -lean_dec(x_57); +lean_object* x_74; +lean_dec(x_70); +lean_dec(x_63); +lean_dec(x_59); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_59); -lean_inc(x_64); -x_72 = l_Lean_Meta_kabstract(x_64, x_59, x_66, x_5, x_6, x_7, x_8, x_69); -if (lean_obj_tag(x_72) == 0) +lean_inc(x_61); +lean_inc(x_66); +x_74 = l_Lean_Meta_kabstract(x_66, x_61, x_68, x_5, x_6, x_7, x_8, x_71); +if (lean_obj_tag(x_74) == 0) { -lean_object* x_73; lean_object* x_74; uint8_t x_75; -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 = l_Lean_Expr_hasLooseBVars(x_73); -if (x_75 == 0) +lean_object* x_75; lean_object* x_76; uint8_t 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_Expr_hasLooseBVars(x_75); +if (x_77 == 0) { -lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; -lean_dec(x_73); -lean_dec(x_70); -lean_dec(x_59); -lean_dec(x_58); -x_76 = l_Lean_indentExpr(x_33); -x_77 = l_Lean_Elab_Term_elabSubst___closed__9; -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_77); -lean_ctor_set(x_78, 1, x_76); -x_79 = l_Lean_Elab_Term_elabSubst___closed__11; +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; uint8_t x_92; +lean_dec(x_75); +lean_dec(x_72); +lean_dec(x_61); +lean_dec(x_60); +x_78 = l_Lean_indentExpr(x_35); +x_79 = l_Lean_Elab_Term_elabSubst___closed__9; x_80 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_80, 0, x_78); -lean_ctor_set(x_80, 1, x_79); -x_81 = l_Lean_indentExpr(x_39); +lean_ctor_set(x_80, 0, x_79); +lean_ctor_set(x_80, 1, x_78); +x_81 = l_Lean_Elab_Term_elabSubst___closed__11; x_82 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_82, 0, x_80); lean_ctor_set(x_82, 1, x_81); -x_83 = l_Lean_Elab_Term_elabSubst___closed__13; +x_83 = l_Lean_indentExpr(x_41); x_84 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_84, 0, x_82); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_indentExpr(x_64); +x_85 = l_Lean_Elab_Term_elabSubst___closed__13; x_86 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_86, 0, x_84); lean_ctor_set(x_86, 1, x_85); -x_87 = l_Lean_Elab_Term_elabCoeFunNotation___closed__4; +x_87 = l_Lean_indentExpr(x_66); x_88 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_88, 0, x_86); lean_ctor_set(x_88, 1, x_87); -x_89 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_74); +x_89 = l_Lean_Elab_Term_elabCoeFunNotation___closed__4; +x_90 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_90, 0, x_88); +lean_ctor_set(x_90, 1, x_89); +x_91 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_90, x_3, x_4, x_5, x_6, x_7, x_8, x_76); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_90 = !lean_is_exclusive(x_89); -if (x_90 == 0) +x_92 = !lean_is_exclusive(x_91); +if (x_92 == 0) { -return x_89; +return x_91; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_89, 0); -x_92 = lean_ctor_get(x_89, 1); -lean_inc(x_92); -lean_inc(x_91); -lean_dec(x_89); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_93 = lean_ctor_get(x_91, 0); +x_94 = lean_ctor_get(x_91, 1); +lean_inc(x_94); +lean_inc(x_93); +lean_dec(x_91); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_93); +lean_ctor_set(x_95, 1, x_94); +return x_95; } } else { -lean_object* x_94; lean_object* x_95; -lean_dec(x_64); -lean_dec(x_39); -x_94 = lean_box(0); -x_95 = l_Lean_Elab_Term_elabSubst___lambda__5(x_70, x_73, x_33, x_58, x_59, x_94, x_3, x_4, x_5, x_6, x_7, x_8, x_74); -return x_95; +lean_object* x_96; lean_object* x_97; +lean_dec(x_66); +lean_dec(x_41); +x_96 = lean_box(0); +x_97 = l_Lean_Elab_Term_elabSubst___lambda__5(x_72, x_75, x_35, x_60, x_61, x_96, x_3, x_4, x_5, x_6, x_7, x_8, x_76); +return x_97; } } else { -uint8_t x_96; -lean_dec(x_70); -lean_dec(x_64); -lean_dec(x_59); -lean_dec(x_58); -lean_dec(x_39); -lean_dec(x_33); +uint8_t x_98; +lean_dec(x_72); +lean_dec(x_66); +lean_dec(x_61); +lean_dec(x_60); +lean_dec(x_41); +lean_dec(x_35); 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_96 = !lean_is_exclusive(x_72); -if (x_96 == 0) +x_98 = !lean_is_exclusive(x_74); +if (x_98 == 0) { -return x_72; +return x_74; } else { -lean_object* x_97; lean_object* x_98; lean_object* x_99; -x_97 = lean_ctor_get(x_72, 0); -x_98 = lean_ctor_get(x_72, 1); -lean_inc(x_98); -lean_inc(x_97); -lean_dec(x_72); -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_99; lean_object* x_100; lean_object* x_101; +x_99 = lean_ctor_get(x_74, 0); +x_100 = lean_ctor_get(x_74, 1); +lean_inc(x_100); +lean_inc(x_99); +lean_dec(x_74); +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; } } } else { -lean_object* x_100; lean_object* x_101; -lean_dec(x_70); -lean_dec(x_64); -lean_dec(x_39); -x_100 = lean_box(0); -x_101 = l_Lean_Elab_Term_elabSubst___lambda__4(x_57, x_61, x_68, x_33, x_58, x_59, x_100, x_3, x_4, x_5, x_6, x_7, x_8, x_69); -lean_dec(x_59); -return x_101; +lean_object* x_102; lean_object* x_103; +lean_dec(x_72); +lean_dec(x_66); +lean_dec(x_41); +x_102 = lean_box(0); +x_103 = l_Lean_Elab_Term_elabSubst___lambda__4(x_59, x_63, x_70, x_35, x_60, x_61, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_71); +lean_dec(x_61); +return x_103; } } else { -uint8_t x_102; -lean_dec(x_64); +uint8_t x_104; +lean_dec(x_66); +lean_dec(x_63); lean_dec(x_61); +lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); -lean_dec(x_57); -lean_dec(x_39); -lean_dec(x_33); +lean_dec(x_41); +lean_dec(x_35); 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_102 = !lean_is_exclusive(x_67); -if (x_102 == 0) +x_104 = !lean_is_exclusive(x_69); +if (x_104 == 0) { -return x_67; +return x_69; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_67, 0); -x_104 = lean_ctor_get(x_67, 1); -lean_inc(x_104); -lean_inc(x_103); -lean_dec(x_67); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -return x_105; +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_69, 0); +x_106 = lean_ctor_get(x_69, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_69); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } } } else { -uint8_t x_106; +uint8_t x_108; +lean_dec(x_63); lean_dec(x_61); +lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); -lean_dec(x_57); -lean_dec(x_39); -lean_dec(x_33); +lean_dec(x_41); +lean_dec(x_35); 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_106 = !lean_is_exclusive(x_63); -if (x_106 == 0) +x_108 = !lean_is_exclusive(x_65); +if (x_108 == 0) { -return x_63; +return x_65; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_63, 0); -x_108 = lean_ctor_get(x_63, 1); -lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_63); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_109 = lean_ctor_get(x_65, 0); +x_110 = lean_ctor_get(x_65, 1); +lean_inc(x_110); +lean_inc(x_109); +lean_dec(x_65); +x_111 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_111, 0, x_109); +lean_ctor_set(x_111, 1, x_110); +return x_111; } } } else { -uint8_t x_110; +uint8_t x_112; +lean_dec(x_61); +lean_dec(x_60); lean_dec(x_59); -lean_dec(x_58); -lean_dec(x_57); -lean_dec(x_39); -lean_dec(x_33); +lean_dec(x_41); +lean_dec(x_35); 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_110 = !lean_is_exclusive(x_60); -if (x_110 == 0) +x_112 = !lean_is_exclusive(x_62); +if (x_112 == 0) { -return x_60; +return x_62; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_60, 0); -x_112 = lean_ctor_get(x_60, 1); -lean_inc(x_112); -lean_inc(x_111); -lean_dec(x_60); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_62, 0); +x_114 = lean_ctor_get(x_62, 1); +lean_inc(x_114); +lean_inc(x_113); +lean_dec(x_62); +x_115 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_115, 0, x_113); +lean_ctor_set(x_115, 1, x_114); +return x_115; } } } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_114 = lean_ctor_get(x_41, 1); -lean_inc(x_114); -lean_dec(x_41); -x_115 = lean_ctor_get(x_54, 0); -lean_inc(x_115); -lean_dec(x_54); -x_116 = lean_ctor_get(x_55, 0); +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; +x_116 = lean_ctor_get(x_43, 1); lean_inc(x_116); -x_117 = lean_ctor_get(x_55, 1); +lean_dec(x_43); +x_117 = lean_ctor_get(x_56, 0); lean_inc(x_117); -lean_dec(x_55); -x_118 = lean_ctor_get(x_11, 0); +lean_dec(x_56); +x_118 = lean_ctor_get(x_57, 0); lean_inc(x_118); +x_119 = lean_ctor_get(x_57, 1); +lean_inc(x_119); +lean_dec(x_57); +x_120 = lean_ctor_get(x_11, 0); +lean_inc(x_120); lean_dec(x_11); -x_119 = lean_box(0); +x_121 = lean_box(0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_117); -lean_inc(x_118); -x_120 = l_Lean_Meta_kabstract(x_118, x_117, x_119, x_5, x_6, x_7, x_8, x_114); -if (lean_obj_tag(x_120) == 0) +lean_inc(x_119); +lean_inc(x_120); +x_122 = l_Lean_Meta_kabstract(x_120, x_119, x_121, x_5, x_6, x_7, x_8, x_116); +if (lean_obj_tag(x_122) == 0) { -lean_object* x_121; lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_121 = lean_ctor_get(x_120, 0); -lean_inc(x_121); -x_122 = lean_ctor_get(x_120, 1); -lean_inc(x_122); -lean_dec(x_120); +lean_object* x_123; lean_object* x_124; lean_object* x_125; uint8_t x_126; +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_ctor_get(x_122, 1); +lean_inc(x_124); +lean_dec(x_122); lean_inc(x_17); lean_inc(x_23); lean_inc(x_1); -lean_inc(x_118); -lean_inc(x_115); -x_123 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__10___boxed), 18, 6); -lean_closure_set(x_123, 0, x_115); -lean_closure_set(x_123, 1, x_118); -lean_closure_set(x_123, 2, x_1); -lean_closure_set(x_123, 3, x_23); -lean_closure_set(x_123, 4, x_17); -lean_closure_set(x_123, 5, x_28); -x_124 = l_Lean_Expr_hasLooseBVars(x_121); -if (x_124 == 0) -{ -lean_object* x_125; -lean_dec(x_121); -lean_dec(x_115); +lean_inc(x_120); +lean_inc(x_117); +x_125 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabSubst___lambda__10___boxed), 18, 6); +lean_closure_set(x_125, 0, x_117); +lean_closure_set(x_125, 1, x_120); +lean_closure_set(x_125, 2, x_1); +lean_closure_set(x_125, 3, x_23); +lean_closure_set(x_125, 4, x_17); +lean_closure_set(x_125, 5, x_28); +x_126 = l_Lean_Expr_hasLooseBVars(x_123); +if (x_126 == 0) +{ +lean_object* x_127; +lean_dec(x_123); +lean_dec(x_117); lean_dec(x_23); lean_dec(x_17); lean_dec(x_1); @@ -25120,139 +25124,139 @@ lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_116); lean_inc(x_118); -x_125 = l_Lean_Meta_kabstract(x_118, x_116, x_119, x_5, x_6, x_7, x_8, x_122); -if (lean_obj_tag(x_125) == 0) +lean_inc(x_120); +x_127 = l_Lean_Meta_kabstract(x_120, x_118, x_121, x_5, x_6, x_7, x_8, x_124); +if (lean_obj_tag(x_127) == 0) { -lean_object* x_126; lean_object* x_127; uint8_t x_128; -x_126 = lean_ctor_get(x_125, 0); -lean_inc(x_126); -x_127 = lean_ctor_get(x_125, 1); -lean_inc(x_127); -lean_dec(x_125); -x_128 = l_Lean_Expr_hasLooseBVars(x_126); -if (x_128 == 0) +lean_object* x_128; lean_object* x_129; uint8_t x_130; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_ctor_get(x_127, 1); +lean_inc(x_129); +lean_dec(x_127); +x_130 = l_Lean_Expr_hasLooseBVars(x_128); +if (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; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; uint8_t x_143; -lean_dec(x_126); -lean_dec(x_123); -lean_dec(x_117); -lean_dec(x_116); -x_129 = l_Lean_indentExpr(x_118); -x_130 = l_Lean_Elab_Term_elabSubst___closed__15; -x_131 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_Elab_Term_elabSubst___closed__17; +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; uint8_t x_145; +lean_dec(x_128); +lean_dec(x_125); +lean_dec(x_119); +lean_dec(x_118); +x_131 = l_Lean_indentExpr(x_120); +x_132 = l_Lean_Elab_Term_elabSubst___closed__15; x_133 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_indentExpr(x_33); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_131); +x_134 = l_Lean_Elab_Term_elabSubst___closed__17; x_135 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_135, 0, x_133); lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Term_elabSubst___closed__19; +x_136 = l_Lean_indentExpr(x_35); x_137 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_137, 0, x_135); lean_ctor_set(x_137, 1, x_136); -x_138 = l_Lean_indentExpr(x_39); +x_138 = l_Lean_Elab_Term_elabSubst___closed__19; x_139 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_139, 0, x_137); lean_ctor_set(x_139, 1, x_138); -x_140 = l_Lean_Elab_Term_elabSubst___closed__21; +x_140 = l_Lean_indentExpr(x_41); x_141 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_141, 0, x_139); lean_ctor_set(x_141, 1, x_140); -x_142 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_141, x_3, x_4, x_5, x_6, x_7, x_8, x_127); +x_142 = l_Lean_Elab_Term_elabSubst___closed__21; +x_143 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_144 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_143, x_3, x_4, x_5, x_6, x_7, x_8, x_129); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_143 = !lean_is_exclusive(x_142); -if (x_143 == 0) +x_145 = !lean_is_exclusive(x_144); +if (x_145 == 0) { -return x_142; +return x_144; } else { -lean_object* x_144; lean_object* x_145; lean_object* x_146; -x_144 = lean_ctor_get(x_142, 0); -x_145 = lean_ctor_get(x_142, 1); -lean_inc(x_145); -lean_inc(x_144); -lean_dec(x_142); -x_146 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_146, 0, x_144); -lean_ctor_set(x_146, 1, x_145); -return x_146; +lean_object* x_146; lean_object* x_147; lean_object* x_148; +x_146 = lean_ctor_get(x_144, 0); +x_147 = lean_ctor_get(x_144, 1); +lean_inc(x_147); +lean_inc(x_146); +lean_dec(x_144); +x_148 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_148, 0, x_146); +lean_ctor_set(x_148, 1, x_147); +return x_148; } } else { -lean_object* x_147; lean_object* x_148; -lean_dec(x_118); -lean_dec(x_39); -x_147 = lean_box(0); -x_148 = l_Lean_Elab_Term_elabSubst___lambda__5(x_123, x_126, x_33, x_116, x_117, x_147, x_3, x_4, x_5, x_6, x_7, x_8, x_127); -return x_148; +lean_object* x_149; lean_object* x_150; +lean_dec(x_120); +lean_dec(x_41); +x_149 = lean_box(0); +x_150 = l_Lean_Elab_Term_elabSubst___lambda__5(x_125, x_128, x_35, x_118, x_119, x_149, x_3, x_4, x_5, x_6, x_7, x_8, x_129); +return x_150; } } else { -uint8_t x_149; -lean_dec(x_123); +uint8_t x_151; +lean_dec(x_125); +lean_dec(x_120); +lean_dec(x_119); lean_dec(x_118); -lean_dec(x_117); -lean_dec(x_116); -lean_dec(x_39); -lean_dec(x_33); +lean_dec(x_41); +lean_dec(x_35); 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_149 = !lean_is_exclusive(x_125); -if (x_149 == 0) +x_151 = !lean_is_exclusive(x_127); +if (x_151 == 0) { -return x_125; +return x_127; } else { -lean_object* x_150; lean_object* x_151; lean_object* x_152; -x_150 = lean_ctor_get(x_125, 0); -x_151 = lean_ctor_get(x_125, 1); -lean_inc(x_151); -lean_inc(x_150); -lean_dec(x_125); -x_152 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_152, 0, x_150); -lean_ctor_set(x_152, 1, x_151); -return x_152; +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_127, 0); +x_153 = lean_ctor_get(x_127, 1); +lean_inc(x_153); +lean_inc(x_152); +lean_dec(x_127); +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_152); +lean_ctor_set(x_154, 1, x_153); +return x_154; } } } else { -lean_object* x_153; lean_object* x_154; -lean_dec(x_123); -lean_dec(x_39); -x_153 = lean_box(0); -x_154 = l_Lean_Elab_Term_elabSubst___lambda__10(x_115, x_118, x_1, x_23, x_17, x_28, x_121, x_33, x_116, x_117, x_153, x_3, x_4, x_5, x_6, x_7, x_8, x_122); -return x_154; +lean_object* x_155; lean_object* x_156; +lean_dec(x_125); +lean_dec(x_41); +x_155 = lean_box(0); +x_156 = l_Lean_Elab_Term_elabSubst___lambda__10(x_117, x_120, x_1, x_23, x_17, x_28, x_123, x_35, x_118, x_119, x_155, x_3, x_4, x_5, x_6, x_7, x_8, x_124); +return x_156; } } else { -uint8_t x_155; +uint8_t x_157; +lean_dec(x_120); +lean_dec(x_119); lean_dec(x_118); lean_dec(x_117); -lean_dec(x_116); -lean_dec(x_115); -lean_dec(x_39); -lean_dec(x_33); +lean_dec(x_41); +lean_dec(x_35); lean_dec(x_23); lean_dec(x_17); lean_dec(x_8); @@ -25262,23 +25266,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_155 = !lean_is_exclusive(x_120); -if (x_155 == 0) +x_157 = !lean_is_exclusive(x_122); +if (x_157 == 0) { -return x_120; +return x_122; } else { -lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_156 = lean_ctor_get(x_120, 0); -x_157 = lean_ctor_get(x_120, 1); -lean_inc(x_157); -lean_inc(x_156); -lean_dec(x_120); -x_158 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -return x_158; +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_122, 0); +x_159 = lean_ctor_get(x_122, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_122); +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; } } } @@ -25286,9 +25290,9 @@ return x_158; } else { -uint8_t x_159; -lean_dec(x_39); -lean_dec(x_33); +uint8_t x_161; +lean_dec(x_41); +lean_dec(x_35); lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25299,30 +25303,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_159 = !lean_is_exclusive(x_41); -if (x_159 == 0) +x_161 = !lean_is_exclusive(x_43); +if (x_161 == 0) { -return x_41; +return x_43; } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; -x_160 = lean_ctor_get(x_41, 0); -x_161 = lean_ctor_get(x_41, 1); -lean_inc(x_161); -lean_inc(x_160); -lean_dec(x_41); -x_162 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_162, 0, x_160); -lean_ctor_set(x_162, 1, x_161); -return x_162; +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_43, 0); +x_163 = lean_ctor_get(x_43, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_43); +x_164 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +return x_164; } } } else { -uint8_t x_163; -lean_dec(x_33); +uint8_t x_165; +lean_dec(x_35); lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25333,29 +25337,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_163 = !lean_is_exclusive(x_35); -if (x_163 == 0) +x_165 = !lean_is_exclusive(x_37); +if (x_165 == 0) { -return x_35; +return x_37; } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_164 = lean_ctor_get(x_35, 0); -x_165 = lean_ctor_get(x_35, 1); -lean_inc(x_165); -lean_inc(x_164); -lean_dec(x_35); -x_166 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_166, 0, x_164); -lean_ctor_set(x_166, 1, x_165); -return x_166; +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_37, 0); +x_167 = lean_ctor_get(x_37, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_37); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +return x_168; } } } else { -uint8_t x_167; +uint8_t x_169; lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25366,29 +25370,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_167 = !lean_is_exclusive(x_32); -if (x_167 == 0) +x_169 = !lean_is_exclusive(x_34); +if (x_169 == 0) { -return x_32; +return x_34; } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_168 = lean_ctor_get(x_32, 0); -x_169 = lean_ctor_get(x_32, 1); -lean_inc(x_169); -lean_inc(x_168); -lean_dec(x_32); -x_170 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_170, 0, x_168); -lean_ctor_set(x_170, 1, x_169); -return x_170; +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_34, 0); +x_171 = lean_ctor_get(x_34, 1); +lean_inc(x_171); +lean_inc(x_170); +lean_dec(x_34); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +return x_172; } } } else { -uint8_t x_171; +uint8_t x_173; lean_dec(x_23); lean_dec(x_17); lean_dec(x_11); @@ -25399,23 +25403,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_171 = !lean_is_exclusive(x_26); -if (x_171 == 0) +x_173 = !lean_is_exclusive(x_26); +if (x_173 == 0) { return x_26; } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; -x_172 = lean_ctor_get(x_26, 0); -x_173 = lean_ctor_get(x_26, 1); -lean_inc(x_173); -lean_inc(x_172); +lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_174 = lean_ctor_get(x_26, 0); +x_175 = lean_ctor_get(x_26, 1); +lean_inc(x_175); +lean_inc(x_174); lean_dec(x_26); -x_174 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_174, 0, x_172); -lean_ctor_set(x_174, 1, x_173); -return x_174; +x_176 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_176, 0, x_174); +lean_ctor_set(x_176, 1, x_175); +return x_176; } } } @@ -25423,7 +25427,7 @@ return x_174; } else { -uint8_t x_175; +uint8_t x_177; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -25431,23 +25435,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_175 = !lean_is_exclusive(x_10); -if (x_175 == 0) +x_177 = !lean_is_exclusive(x_10); +if (x_177 == 0) { return x_10; } else { -lean_object* x_176; lean_object* x_177; lean_object* x_178; -x_176 = lean_ctor_get(x_10, 0); -x_177 = lean_ctor_get(x_10, 1); -lean_inc(x_177); -lean_inc(x_176); +lean_object* x_178; lean_object* x_179; lean_object* x_180; +x_178 = lean_ctor_get(x_10, 0); +x_179 = lean_ctor_get(x_10, 1); +lean_inc(x_179); +lean_inc(x_178); lean_dec(x_10); -x_178 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_178, 0, x_176); -lean_ctor_set(x_178, 1, x_177); -return x_178; +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; } } } @@ -26640,7 +26644,7 @@ lean_inc(x_20); lean_dec(x_18); lean_inc(x_3); lean_inc(x_19); -x_21 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_17); +x_21 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_19, x_3, x_4, x_5, x_6, x_7, x_8, x_17); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; diff --git a/stage0/stdlib/Lean/Elab/BuiltinTerm.c b/stage0/stdlib/Lean/Elab/BuiltinTerm.c index 24d0ab77c467..74e7f053f364 100644 --- a/stage0/stdlib/Lean/Elab/BuiltinTerm.c +++ b/stage0/stdlib/Lean/Elab/BuiltinTerm.c @@ -564,7 +564,6 @@ lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTypeOf___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabScientificLit_declRange___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_elabCompletion___closed__1; -lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabTypeStx(lean_object*); LEAN_EXPORT lean_object* l_Lean_activateScoped___at_Lean_Elab_Term_elabOpen___spec__31___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -970,6 +969,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabDeclName(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabSort_declRange___closed__6; +lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabOmission___closed__2; static lean_object* l_Lean_Elab_Term_elabDeclName___lambda__1___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_elabProp___closed__2; @@ -4765,7 +4765,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabOmission(lean_object* x_1, lean_ob { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; x_10 = l_Lean_Elab_Term_elabOmission___closed__2; -x_11 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); x_12 = lean_ctor_get(x_11, 1); lean_inc(x_12); lean_dec(x_11); diff --git a/stage0/stdlib/Lean/Elab/CheckTactic.c b/stage0/stdlib/Lean/Elab/CheckTactic.c index c11d99934e80..a82568bf30f0 100644 --- a/stage0/stdlib/Lean/Elab/CheckTactic.c +++ b/stage0/stdlib/Lean/Elab/CheckTactic.c @@ -51,7 +51,6 @@ static lean_object* l_Lean_Elab_CheckTactic_elabCheckTacticFailure___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_elabCheckTacticFailure___closed__2; static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_elabCheckTacticFailure_declRange___closed__6; -lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_CheckTactic_expandCheckSimp(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_expandCheckSimpFailure_declRange___closed__7; @@ -100,6 +99,7 @@ lean_object* l_Lean_Syntax_getArg(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*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_elabCheckTactic_declRange___closed__2; static lean_object* l_Lean_Elab_CheckTactic_expandCheckSimp___closed__2; +lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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_Elab_CheckTactic_elabCheckTactic(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_CheckTactic_elabCheckTacticFailure_declRange___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_CheckTactic_elabCheckTacticFailure(lean_object*, lean_object*, lean_object*, lean_object*); @@ -610,7 +610,7 @@ x_48 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; 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_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_4, x_49, x_6, x_7, x_8, x_9, x_10, x_11, x_40); +x_50 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_4, x_49, x_6, x_7, x_8, x_9, x_10, x_11, x_40); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); @@ -724,7 +724,7 @@ x_83 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; x_84 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_84, 0, x_82); lean_ctor_set(x_84, 1, x_83); -x_85 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_4, x_84, x_6, x_7, x_8, x_9, x_10, x_11, x_75); +x_85 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_4, x_84, x_6, x_7, x_8, x_9, x_10, x_11, x_75); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); @@ -878,7 +878,7 @@ x_124 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__8; x_125 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_125, 0, x_123); lean_ctor_set(x_125, 1, x_124); -x_126 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_4, x_125, x_6, x_7, x_8, x_9, x_10, x_11, x_116); +x_126 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_4, x_125, x_6, x_7, x_8, x_9, x_10, x_11, x_116); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); @@ -1041,7 +1041,7 @@ x_151 = l_Lean_Elab_CheckTactic_elabCheckTactic___lambda__2___closed__12; x_152 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_152, 0, x_150); lean_ctor_set(x_152, 1, x_151); -x_153 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_4, x_152, x_6, x_7, x_8, x_9, x_10, x_11, x_143); +x_153 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_4, x_152, x_6, x_7, x_8, x_9, x_10, x_11, x_143); lean_dec(x_11); lean_dec(x_9); lean_dec(x_8); @@ -2031,7 +2031,7 @@ x_47 = l_Lean_Elab_CheckTactic_elabCheckTacticFailure___lambda__2___closed__4; x_48 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_48, 0, x_46); lean_ctor_set(x_48, 1, x_47); -x_49 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_3, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_35); +x_49 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_3, x_48, x_5, x_6, x_7, x_8, x_9, x_10, x_35); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -2106,7 +2106,7 @@ lean_ctor_set(x_71, 1, x_62); x_72 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_72, 0, x_71); lean_ctor_set(x_72, 1, x_60); -x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_3, x_72, x_5, x_6, x_7, x_8, x_9, x_10, x_57); +x_73 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_3, x_72, x_5, x_6, x_7, x_8, x_9, x_10, x_57); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); @@ -2213,7 +2213,7 @@ lean_inc(x_92); x_93 = lean_ctor_get(x_91, 1); lean_inc(x_93); lean_dec(x_91); -x_94 = l_Lean_throwErrorAt___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_ensureAtomicBinderName___spec__1(x_3, x_92, x_5, x_6, x_7, x_8, x_9, x_10, x_93); +x_94 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_3, x_92, x_5, x_6, x_7, x_8, x_9, x_10, x_93); lean_dec(x_10); lean_dec(x_8); lean_dec(x_7); diff --git a/stage0/stdlib/Lean/Elab/Command.c b/stage0/stdlib/Lean/Elab/Command.c index b868b75811e9..29ac853e6f9c 100644 --- a/stage0/stdlib/Lean/Elab/Command.c +++ b/stage0/stdlib/Lean/Elab/Command.c @@ -42,9 +42,7 @@ lean_object* l_Lean_MapDeclarationExtension_insert___rarg(lean_object*, lean_obj lean_object* l_instBEqOfDecidableEq___rarg(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; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10; extern lean_object* l_Lean_diagExt; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13; extern lean_object* l_Lean_profiler; lean_object* l_Lean_Core_getMaxHeartbeats(lean_object*); uint8_t l_Lean_Elab_isAbortExceptionId(lean_object*); @@ -72,6 +70,7 @@ static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabComma static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__6; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__1___closed__2; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__12; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__2; lean_object* lean_private_to_user_name(lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___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_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -92,17 +91,22 @@ lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_opts___default; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_checkIfShadowingStructureField___at_Lean_Elab_Command_expandDeclId___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandDeclId___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__1; static lean_object* l_Lean_Elab_Command_instMonadInfoTreeCommandElabM___closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_Command_instInhabitedCommandElabM___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_getLevelNames___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___rarg___closed__1; lean_object* l_Lean_ConstantInfo_levelParams(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7; lean_object* l_Lean_indentD(lean_object*); double lean_float_div(double, double); LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); @@ -203,7 +207,6 @@ lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Command_expandDeclId___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_getScopes___rarg(lean_object*, lean_object*); 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_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Command_withMacroExpansion___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed__3; static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; @@ -224,7 +227,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftCoreM___rarg___boxed(lean_objec LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__8(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2; lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -236,6 +238,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_expandDeclId___at_Lean_Elab_Command_expandD static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__9; 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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__11(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__4; extern lean_object* l_Lean_LocalContext_empty; @@ -273,6 +276,7 @@ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCom static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__7; LEAN_EXPORT lean_object* l_Lean_liftCommandElabM___rarg___boxed(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_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8; LEAN_EXPORT lean_object* l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext(lean_object*, lean_object*); @@ -281,6 +285,7 @@ static lean_object* l_Lean_logAt___at_Lean_Elab_Command_runLinters___spec__2___c LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Command_liftTermElabM___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_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); @@ -291,6 +296,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_addCompletionInfo___at_Lean_withSetOptionIn___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at_Lean_Elab_Command_elabCommandTopLevel___spec__4(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_3250____closed__2; static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__18___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); @@ -324,11 +330,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_addT LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___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__10___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878_(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9; 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, 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_Elab_Command_instMonadRecDepthCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908_(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_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; @@ -341,6 +349,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandT LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___rarg(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_getVarDecls___boxed(lean_object*); static lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___closed__2; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1; static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__24; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM; @@ -351,6 +360,7 @@ LEAN_EXPORT lean_object* l_Lean_mkHashMap___at___private_Lean_Elab_Command_0__Le 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; static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439_(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; @@ -370,7 +380,7 @@ static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_el LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__10(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_nextMacroScope___default; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Command_elabCommand___spec__17___boxed(lean_object*, lean_object*); static lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -435,7 +445,6 @@ LEAN_EXPORT lean_object* l_Lean_liftCommandElabM(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__22; 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_object* l_Lean_Meta_MetaM_run___rarg(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_547____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___lambda__3___boxed(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*); @@ -457,7 +466,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_scopes___default; static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___lambda__5___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17; lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_State_ngen___default___closed__3; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__4; @@ -479,6 +487,7 @@ 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; static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_State_ngen___default; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at_Lean_Elab_Command_elabCommandTopLevel___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -501,7 +510,6 @@ 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_Elab_Command_getMainModule___rarg___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__9; lean_object* l_Lean_Elab_expandDeclIdCore(lean_object*); static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__7; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Elab_Command_runLinters___spec__15___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -524,7 +532,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption_setOption___at_Lean_withSetOptionIn___spec__4___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Elab_Command_elabCommand___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__7(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__1; @@ -556,6 +563,7 @@ static lean_object* l_Lean_Elab_Command_instMonadQuotationCommandElabM___closed_ static lean_object* l_Lean_liftCommandElabM___rarg___closed__2; static lean_object* l_Lean_Elab_Command_mkState___closed__1; uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3; 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; @@ -564,7 +572,6 @@ static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_currNamespace___default; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runLinters___spec__9___lambda__1___closed__2; extern lean_object* l_Lean_diagnostics; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14; static lean_object* l_Lean_Elab_Command_elabCommandTopLevel___closed__1; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__1; static lean_object* l_Lean_Elab_Command_runLinters___lambda__1___closed__1; @@ -619,7 +626,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_withMacroExpansion___rarg___lambda_ static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__4; static lean_object* l_Lean_Elab_Command_elabCommand___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_Scope_varDecls___default; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__6; static lean_object* l_Lean_Elab_Command_modifyScope___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -627,15 +633,17 @@ lean_object* l_Lean_Elab_Term_withAutoBoundImplicit___rarg(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__5(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lean_object*); double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_qsort_sort___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__19___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at_Lean_Elab_Command_runLinters___spec__5(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_elabCommand___lambda__4___closed__5; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11; static lean_object* l_Lean_addMessageContextPartial___at_Lean_Elab_Command_instAddMessageContextCommandElabM___spec__1___closed__5; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Command_liftTermElabM___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*); LEAN_EXPORT lean_object* l___auto____x40_Lean_Elab_Command___hyg_403_; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_elabCommand___spec__15(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__15; @@ -645,7 +653,6 @@ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__19; static double l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadTraceCommandElabM___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___closed__1; static lean_object* l_Lean_Elab_Command_State_messages___default___closed__1; @@ -653,7 +660,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___s LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Command_expandDeclId___spec__12(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___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_elabCommand___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_515_(lean_object*); 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; @@ -705,7 +711,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_getMainModule(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_Command_elabCommand___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Command_expandDeclId___spec__7___closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2; static lean_object* l_Lean_addDocString___at_Lean_Elab_Command_expandDeclId___spec__14___closed__1; static lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___closed__2; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__8; @@ -739,7 +744,6 @@ extern lean_object* l_Lean_Expr_instHashable; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDeclId___spec__2___closed__2; static lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__2___closed__2; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadEnvCommandElabM___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instAddErrorMessageContextCommandElabM(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -747,7 +751,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadCommandElabM___lambda__2(l static lean_object* l_Lean_Elab_Command_instMonadResolveNameCommandElabM___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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runLinters___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__18(lean_object*, size_t, size_t, lean_object*); @@ -771,7 +775,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters_ static lean_object* l___auto____x40_Lean_Elab_Command___hyg_403____closed__28; 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5; static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__14; lean_object* l_Lean_Kernel_resetDiag___lambda__1(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__19___closed__2; @@ -792,17 +795,13 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Command_expandDeclId__ 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*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3; static lean_object* l_List_foldl___at_Lean_Elab_Command_instAddErrorMessageContextCommandElabM___spec__2___closed__2; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__3; static lean_object* l_Lean_Elab_Command_instMonadCommandElabM___closed__7; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11; extern lean_object* l_Lean_firstFrontendMacroScope; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabCommand___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instInhabitedScope; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -816,6 +815,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_mkDeclName___at_Lean_Elab_Command_expandDec static lean_object* l_Lean_Elab_Command_liftTermElabM___rarg___closed__17; size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadLogCommandElabM___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17; LEAN_EXPORT uint8_t l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___elambda__1(lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_Command_Scope_isNoncomputable___default; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -834,13 +834,11 @@ 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; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1; lean_object* lean_io_error_to_string(lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_getEntries___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabSetOption___at_Lean_withSetOptionIn___spec__1___closed__6; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3; static lean_object* l_Lean_Elab_logException___at_Lean_Elab_Command_runLinters___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__20___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_mkMetaContext; @@ -854,6 +852,7 @@ static lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_liftCommandElabM___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getInfoTrees___at_Lean_Elab_Command_elabCommandTopLevel___spec__1___boxed(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__4___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_trace_profiler_output; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__4; @@ -888,7 +887,7 @@ static lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab 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; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Command_0__Lean_Elab_Command_addTraceAsMessagesCore___spec__20___closed__5; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12; +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_485_(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Command_expandDeclId___spec__5___closed__4; @@ -901,8 +900,6 @@ 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_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Command_runLinters___spec__13(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_expandDeclId___spec__19(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; 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*); static lean_object* l_Lean_Elab_Command_instMonadMacroAdapterCommandElabM___closed__2; LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__16(lean_object*, lean_object*, lean_object*, lean_object*); @@ -916,8 +913,8 @@ 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_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean_Elab_Command_addUnivLevel___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469_(lean_object*); static lean_object* l_Lean_Elab_Command_State_infoState___default___closed__4; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15; static lean_object* l_Lean_Elab_Command_getBracketedBinderIds___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___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___private_Lean_Elab_Command_0__Lean_Elab_Command_mkTermContext___boxed(lean_object*, lean_object*); @@ -935,6 +932,8 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lam 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*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_instMonadRecDepthCommandElabM___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_runCore(lean_object*); @@ -951,6 +950,7 @@ 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; LEAN_EXPORT lean_object* l___private_Lean_Elab_Command_0__Lean_Elab_Command_getVarDecls(lean_object*); +static lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2; uint8_t l_Array_isEmpty___rarg(lean_object*); static lean_object* _init_l_Lean_Elab_Command_Scope_opts___default() { _start: @@ -1796,7 +1796,7 @@ lean_ctor_set(x_17, 7, x_16); return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_515_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_485_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -1822,7 +1822,7 @@ return x_7; } } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1() { _start: { lean_object* x_1; @@ -1830,7 +1830,7 @@ x_1 = lean_mk_string_from_bytes("Elab", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2() { _start: { lean_object* x_1; @@ -1838,17 +1838,17 @@ x_1 = lean_mk_string_from_bytes("lint", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -1858,17 +1858,17 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6() { _start: { lean_object* x_1; @@ -1876,17 +1876,17 @@ x_1 = lean_mk_string_from_bytes("Command", 7); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__7() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8() { _start: { lean_object* x_1; @@ -1894,17 +1894,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_Command___hyg_547____closed__9() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__7; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10() { _start: { lean_object* x_1; @@ -1912,47 +1912,47 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__9; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11; x_2 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15() { _start: { lean_object* x_1; @@ -1960,33 +1960,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_Command___hyg_547____closed__16() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16; -x_2 = lean_unsigned_to_nat(547u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16; +x_2 = lean_unsigned_to_nat(517u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -12569,7 +12569,7 @@ lean_dec(x_8); x_11 = lean_ctor_get(x_10, 1); lean_inc(x_11); lean_dec(x_10); -x_12 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3; +x_12 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3; x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runLinters___lambda__2), 5, 2); lean_closure_set(x_13, 0, x_1); lean_closure_set(x_13, 1, x_12); @@ -13225,7 +13225,7 @@ static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___clo lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__1; x_2 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } @@ -13243,8 +13243,8 @@ static lean_object* _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___clo { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_4 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__6; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -13271,7 +13271,7 @@ x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_3, x_4, x_5, x_6, x_7, x_1, x_2); return x_8; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1() { _start: { lean_object* x_1; @@ -13279,23 +13279,23 @@ x_1 = lean_mk_string_from_bytes("commandElabAttribute", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____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___auto____x40_Lean_Elab_Command___hyg_403____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__2; x_3 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe(x_2, x_1); return x_3; } @@ -15264,33 +15264,33 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; x_2 = l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16; -x_2 = lean_unsigned_to_nat(2908u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16; +x_2 = lean_unsigned_to_nat(2878u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -18404,7 +18404,7 @@ static lean_object* _init_l_Lean_Elab_Command_elabCommand___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___auto____x40_Lean_Elab_Command___hyg_403____closed__1; x_2 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_4 = l_Lean_Elab_Command_elabCommand___closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -18460,7 +18460,7 @@ lean_inc(x_1); x_13 = l_Lean_Syntax_getKind(x_1); x_14 = 1; x_15 = l_Lean_Name_toString(x_13, x_14); -x_16 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1; +x_16 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1; x_17 = lean_box(x_14); x_18 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___boxed), 8, 5); lean_closure_set(x_18, 0, x_16); @@ -18482,7 +18482,7 @@ x_22 = l_Lean_Syntax_getArg(x_1, x_21); x_23 = l_Lean_Syntax_getKind(x_22); x_24 = 1; x_25 = l_Lean_Name_toString(x_23, x_24); -x_26 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1; +x_26 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1; x_27 = lean_box(x_24); x_28 = lean_alloc_closure((void*)(l_Lean_withTraceNode___at_Lean_Elab_Command_runLinters___spec__10___boxed), 8, 5); lean_closure_set(x_28, 0, x_26); @@ -18755,7 +18755,7 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1() { _start: { lean_object* x_1; @@ -18763,33 +18763,33 @@ x_1 = lean_mk_string_from_bytes("input", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3() { +static lean_object* _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16; -x_2 = lean_unsigned_to_nat(3280u); +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16; +x_2 = lean_unsigned_to_nat(3250u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2; +x_2 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3; +x_4 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -19217,7 +19217,7 @@ static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCo _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1; +x_1 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1; x_2 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; @@ -24848,121 +24848,122 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Command_runTermElabM___rarg___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -uint8_t x_11; lean_object* x_12; -x_11 = 0; +uint8_t x_11; uint8_t x_12; lean_object* x_13; +x_11 = 1; +x_12 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_12 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_11, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_12) == 0) +x_13 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_13) == 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; size_t 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_13 = lean_ctor_get(x_12, 1); -lean_inc(x_13); -lean_dec(x_12); -x_14 = lean_box(0); -x_15 = lean_array_get_size(x_3); -x_16 = lean_unsigned_to_nat(0u); +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; uint8_t x_28; +x_14 = lean_ctor_get(x_13, 1); +lean_inc(x_14); +lean_dec(x_13); +x_15 = lean_box(0); +x_16 = lean_array_get_size(x_3); +x_17 = lean_unsigned_to_nat(0u); lean_inc(x_3); -x_17 = l_Array_toSubarray___rarg(x_3, x_16, x_15); -x_18 = lean_ctor_get(x_1, 6); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_17); -lean_ctor_set(x_19, 1, x_14); -x_20 = lean_array_get_size(x_18); -x_21 = lean_usize_of_nat(x_20); -lean_dec(x_20); -x_22 = 0; -x_23 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(x_18, x_21, x_22, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_13); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +x_18 = l_Array_toSubarray___rarg(x_3, x_17, x_16); +x_19 = lean_ctor_get(x_1, 6); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_18); +lean_ctor_set(x_20, 1, x_15); +x_21 = lean_array_get_size(x_19); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +x_23 = 0; +x_24 = l_Array_forInUnsafe_loop___at_Lean_Elab_Command_runTermElabM___spec__1(x_19, x_22, x_23, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); x_26 = lean_ctor_get(x_24, 1); lean_inc(x_26); lean_dec(x_24); -x_27 = !lean_is_exclusive(x_4); -if (x_27 == 0) +x_27 = lean_ctor_get(x_25, 1); +lean_inc(x_27); +lean_dec(x_25); +x_28 = !lean_is_exclusive(x_4); +if (x_28 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_28 = lean_ctor_get(x_4, 6); -lean_dec(x_28); -lean_ctor_set(x_4, 6, x_26); -x_29 = l_Lean_Core_resetMessageLog(x_8, x_9, x_25); -x_30 = lean_ctor_get(x_29, 1); -lean_inc(x_30); +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_4, 6); lean_dec(x_29); -x_31 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runTermElabM___rarg___lambda__1___boxed), 10, 1); -lean_closure_set(x_31, 0, x_2); -x_32 = l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed__1; -x_33 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_32, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_30); -return x_33; +lean_ctor_set(x_4, 6, x_27); +x_30 = l_Lean_Core_resetMessageLog(x_8, x_9, x_26); +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runTermElabM___rarg___lambda__1___boxed), 10, 1); +lean_closure_set(x_32, 0, x_2); +x_33 = l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed__1; +x_34 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_33, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_31); +return x_34; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; lean_object* x_48; uint8_t x_49; uint8_t x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_34 = lean_ctor_get(x_4, 0); -x_35 = lean_ctor_get(x_4, 1); -x_36 = lean_ctor_get(x_4, 2); -x_37 = lean_ctor_get_uint8(x_4, sizeof(void*)*8); -x_38 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 1); -x_39 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 2); -x_40 = lean_ctor_get(x_4, 3); -x_41 = lean_ctor_get(x_4, 4); -x_42 = lean_ctor_get(x_4, 5); -x_43 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 3); -x_44 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 4); -x_45 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 5); -x_46 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 6); -x_47 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 7); -x_48 = lean_ctor_get(x_4, 7); -x_49 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 8); -x_50 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 9); -lean_inc(x_48); +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; lean_object* x_49; uint8_t x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_35 = lean_ctor_get(x_4, 0); +x_36 = lean_ctor_get(x_4, 1); +x_37 = lean_ctor_get(x_4, 2); +x_38 = lean_ctor_get_uint8(x_4, sizeof(void*)*8); +x_39 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 1); +x_40 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 2); +x_41 = lean_ctor_get(x_4, 3); +x_42 = lean_ctor_get(x_4, 4); +x_43 = lean_ctor_get(x_4, 5); +x_44 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 3); +x_45 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 4); +x_46 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 5); +x_47 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 6); +x_48 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 7); +x_49 = lean_ctor_get(x_4, 7); +x_50 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 8); +x_51 = lean_ctor_get_uint8(x_4, sizeof(void*)*8 + 9); +lean_inc(x_49); +lean_inc(x_43); lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); +lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); lean_dec(x_4); -x_51 = lean_alloc_ctor(0, 8, 10); -lean_ctor_set(x_51, 0, x_34); -lean_ctor_set(x_51, 1, x_35); -lean_ctor_set(x_51, 2, x_36); -lean_ctor_set(x_51, 3, x_40); -lean_ctor_set(x_51, 4, x_41); -lean_ctor_set(x_51, 5, x_42); -lean_ctor_set(x_51, 6, x_26); -lean_ctor_set(x_51, 7, x_48); -lean_ctor_set_uint8(x_51, sizeof(void*)*8, x_37); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 1, x_38); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 2, x_39); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 3, x_43); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 4, x_44); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 5, x_45); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 6, x_46); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 7, x_47); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 8, x_49); -lean_ctor_set_uint8(x_51, sizeof(void*)*8 + 9, x_50); -x_52 = l_Lean_Core_resetMessageLog(x_8, x_9, x_25); -x_53 = lean_ctor_get(x_52, 1); -lean_inc(x_53); -lean_dec(x_52); -x_54 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runTermElabM___rarg___lambda__1___boxed), 10, 1); -lean_closure_set(x_54, 0, x_2); -x_55 = l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed__1; -x_56 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_55, x_54, x_51, x_5, x_6, x_7, x_8, x_9, x_53); -return x_56; +x_52 = lean_alloc_ctor(0, 8, 10); +lean_ctor_set(x_52, 0, x_35); +lean_ctor_set(x_52, 1, x_36); +lean_ctor_set(x_52, 2, x_37); +lean_ctor_set(x_52, 3, x_41); +lean_ctor_set(x_52, 4, x_42); +lean_ctor_set(x_52, 5, x_43); +lean_ctor_set(x_52, 6, x_27); +lean_ctor_set(x_52, 7, x_49); +lean_ctor_set_uint8(x_52, sizeof(void*)*8, x_38); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 1, x_39); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 2, x_40); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 3, x_44); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 4, x_45); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 5, x_46); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 6, x_47); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 7, x_48); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 8, x_50); +lean_ctor_set_uint8(x_52, sizeof(void*)*8 + 9, x_51); +x_53 = l_Lean_Core_resetMessageLog(x_8, x_9, x_26); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = lean_alloc_closure((void*)(l_Lean_Elab_Command_runTermElabM___rarg___lambda__1___boxed), 10, 1); +lean_closure_set(x_55, 0, x_2); +x_56 = l_Lean_Elab_Command_runTermElabM___rarg___lambda__2___closed__1; +x_57 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_56, x_55, x_52, x_5, x_6, x_7, x_8, x_9, x_54); +return x_57; } } else { -uint8_t x_57; +uint8_t x_58; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -24971,23 +24972,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_57 = !lean_is_exclusive(x_12); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_13); +if (x_58 == 0) { -return x_12; +return x_13; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_12, 0); -x_59 = lean_ctor_get(x_12, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_13, 0); +x_60 = lean_ctor_get(x_13, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_12); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_13); +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; } } } @@ -31261,7 +31262,7 @@ static lean_object* _init_l_Lean_withSetOptionIn___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___auto____x40_Lean_Elab_Command___hyg_403____closed__1; x_2 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_4 = l_Lean_withSetOptionIn___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -31281,7 +31282,7 @@ static lean_object* _init_l_Lean_withSetOptionIn___closed__4() { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__1; x_2 = l___auto____x40_Lean_Elab_Command___hyg_403____closed__2; -x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6; +x_3 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6; x_4 = l_Lean_withSetOptionIn___closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; @@ -31586,46 +31587,46 @@ l_Lean_Elab_Command_instMonadCommandElabM = _init_l_Lean_Elab_Command_instMonadC lean_mark_persistent(l_Lean_Elab_Command_instMonadCommandElabM); l_Lean_Elab_Command_mkState___closed__1 = _init_l_Lean_Elab_Command_mkState___closed__1(); lean_mark_persistent(l_Lean_Elab_Command_mkState___closed__1); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_515_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_485_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_lintersRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_lintersRef); lean_dec_ref(res); -}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__3); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__4); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__5); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__6); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__7(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__7); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__9(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__9); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__10); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__11); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__12); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__13); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__14); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__15); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__16); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547____closed__17); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_547_(lean_io_mk_world()); +}l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__3); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__4); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__5); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__6); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__7); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__8); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__9); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__10); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__11); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__12); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__13); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__14); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__15); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__16); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517____closed__17); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_517_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Elab_Command_addLinter___closed__1 = _init_l_Lean_Elab_Command_addLinter___closed__1(); @@ -31809,11 +31810,11 @@ l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__7); l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8 = _init_l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8(); lean_mark_persistent(l_Lean_Elab_Command_mkCommandElabAttributeUnsafe___closed__8); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469____closed__2); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2469_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439____closed__2); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2439_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Command_commandElabAttribute = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Command_commandElabAttribute); @@ -31846,11 +31847,11 @@ l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4 = _init_l_Lean_Ela lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM___closed__4); l_Lean_Elab_Command_instMonadRecDepthCommandElabM = _init_l_Lean_Elab_Command_instMonadRecDepthCommandElabM(); lean_mark_persistent(l_Lean_Elab_Command_instMonadRecDepthCommandElabM); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908____closed__2); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2908_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878____closed__2); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_2878_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Command_elabCommand___spec__1___closed__1(); @@ -31883,13 +31884,13 @@ l_Lean_Elab_Command_elabCommand___closed__6 = _init_l_Lean_Elab_Command_elabComm lean_mark_persistent(l_Lean_Elab_Command_elabCommand___closed__6); l_Lean_Elab_Command_elabCommand___boxed__const__1 = _init_l_Lean_Elab_Command_elabCommand___boxed__const__1(); lean_mark_persistent(l_Lean_Elab_Command_elabCommand___boxed__const__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__1); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__2); -l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3(); -lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280____closed__3); -if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3280_(lean_io_mk_world()); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__1); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__2 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__2(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__2); +l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3 = _init_l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3(); +lean_mark_persistent(l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250____closed__3); +if (builtin) {res = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Command___hyg_3250_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Command_elabCommandTopLevel___spec__6___closed__1(); diff --git a/stage0/stdlib/Lean/Elab/Declaration.c b/stage0/stdlib/Lean/Elab/Declaration.c index 4506fd872b31..8519e86b3d2f 100644 --- a/stage0/stdlib/Lean/Elab/Declaration.c +++ b/stage0/stdlib/Lean/Elab/Declaration.c @@ -2328,170 +2328,171 @@ lean_inc(x_10); x_21 = l_Lean_Elab_Term_elabType(x_3, x_10, x_11, x_12, x_13, x_14, x_15, x_20); if (lean_obj_tag(x_21) == 0) { -lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; +lean_object* x_22; lean_object* x_23; uint8_t x_24; uint8_t 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 = 0; +x_24 = 1; +x_25 = 0; 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_25 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_24, x_24, x_10, x_11, x_12, x_13, x_14, x_15, x_23); -if (lean_obj_tag(x_25) == 0) +x_26 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_24, x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_23); +if (lean_obj_tag(x_26) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; uint8_t x_31; lean_object* x_32; -x_26 = lean_ctor_get(x_25, 1); -lean_inc(x_26); -lean_dec(x_25); -x_27 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_26); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; +x_27 = lean_ctor_get(x_26, 1); +lean_inc(x_27); +lean_dec(x_26); +x_28 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_27); +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = 1; +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); x_31 = 1; -x_32 = l_Lean_Meta_mkForallFVars(x_9, x_28, x_24, x_30, x_31, x_12, x_13, x_14, x_15, x_29); -if (lean_obj_tag(x_32) == 0) +x_32 = 1; +x_33 = l_Lean_Meta_mkForallFVars(x_9, x_29, x_25, x_31, x_32, x_12, x_13, x_14, x_15, x_30); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Meta_mkForallFVars(x_4, x_33, x_30, x_30, x_31, x_12, x_13, x_14, x_15, x_34); -if (lean_obj_tag(x_35) == 0) +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Meta_mkForallFVars(x_4, x_34, x_31, x_31, x_32, x_12, x_13, x_14, x_15, x_35); +if (lean_obj_tag(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; -x_36 = lean_ctor_get(x_35, 0); -lean_inc(x_36); -x_37 = lean_ctor_get(x_35, 1); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_37 = lean_ctor_get(x_36, 0); lean_inc(x_37); -lean_dec(x_35); -x_38 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__1; -x_39 = l_Lean_Elab_Term_levelMVarToParam(x_36, x_38, x_10, x_11, x_12, x_13, x_14, x_15, x_37); -x_40 = lean_ctor_get(x_39, 0); -lean_inc(x_40); -x_41 = lean_ctor_get(x_39, 1); +x_38 = lean_ctor_get(x_36, 1); +lean_inc(x_38); +lean_dec(x_36); +x_39 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__1; +x_40 = l_Lean_Elab_Term_levelMVarToParam(x_37, x_39, x_10, x_11, x_12, x_13, x_14, x_15, x_38); +x_41 = lean_ctor_get(x_40, 0); lean_inc(x_41); -lean_dec(x_39); -x_42 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__4; -lean_inc(x_40); -x_43 = l_Lean_CollectLevelParams_main(x_40, x_42); -x_44 = lean_ctor_get(x_43, 2); -lean_inc(x_44); -lean_dec(x_43); -x_45 = l_Lean_Elab_sortDeclLevelParams(x_5, x_6, x_44); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_42 = lean_ctor_get(x_40, 1); +lean_inc(x_42); lean_dec(x_40); +x_43 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__4; +lean_inc(x_41); +x_44 = l_Lean_CollectLevelParams_main(x_41, x_43); +x_45 = lean_ctor_get(x_44, 2); +lean_inc(x_45); +lean_dec(x_44); +x_46 = l_Lean_Elab_sortDeclLevelParams(x_5, x_6, x_45); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +lean_dec(x_41); lean_dec(x_17); lean_dec(x_8); lean_dec(x_2); lean_dec(x_1); -x_46 = lean_ctor_get(x_45, 0); -lean_inc(x_46); -lean_dec(x_45); -x_47 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(0, 1, 0); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +lean_dec(x_46); +x_48 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_48, 0, x_47); -x_49 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_7, x_48, x_10, x_11, x_12, x_13, x_14, x_15, x_41); +x_49 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_49, 0, x_48); +x_50 = l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16(x_7, x_49, x_10, x_11, x_12, x_13, x_14, x_15, x_42); lean_dec(x_15); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_7); -return x_49; +return x_50; } else { -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_object* x_60; uint8_t x_61; +lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; uint8_t x_62; lean_dec(x_7); -x_50 = lean_ctor_get(x_45, 0); -lean_inc(x_50); -lean_dec(x_45); -x_51 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_40, x_10, x_11, x_12, x_13, x_14, x_15, x_41); -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_ctor_get(x_51, 1); +x_51 = lean_ctor_get(x_46, 0); +lean_inc(x_51); +lean_dec(x_46); +x_52 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_41, x_10, x_11, x_12, x_13, x_14, x_15, x_42); +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_53); -lean_dec(x_51); -lean_inc(x_52); lean_inc(x_2); -x_54 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_54, 0, x_2); -lean_ctor_set(x_54, 1, x_50); -lean_ctor_set(x_54, 2, x_52); -x_55 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 3); +x_55 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_55, 0, x_2); +lean_ctor_set(x_55, 1, x_51); +lean_ctor_set(x_55, 2, x_53); +x_56 = lean_ctor_get_uint8(x_1, sizeof(void*)*2 + 3); lean_dec(x_1); -x_56 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_56, 0, x_54); -lean_ctor_set_uint8(x_56, sizeof(void*)*1, x_55); -x_57 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_57, 0, x_56); -x_58 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__6; -x_59 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_58, x_10, x_11, x_12, x_13, x_14, x_15, x_53); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_unbox(x_60); -lean_dec(x_60); -if (x_61 == 0) +x_57 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set_uint8(x_57, sizeof(void*)*1, x_56); +x_58 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_59 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__6; +x_60 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_59, x_10, x_11, x_12, x_13, x_14, x_15, x_54); +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) { -lean_object* x_62; lean_object* x_63; lean_object* x_64; -lean_dec(x_52); -x_62 = lean_ctor_get(x_59, 1); -lean_inc(x_62); -lean_dec(x_59); -x_63 = lean_box(0); -x_64 = l_Lean_Elab_Command_elabAxiom___lambda__4(x_57, x_2, x_8, x_17, x_63, x_10, x_11, x_12, x_13, x_14, x_15, x_62); -return x_64; +lean_object* x_63; lean_object* x_64; lean_object* x_65; +lean_dec(x_53); +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_Elab_Command_elabAxiom___lambda__4(x_58, x_2, x_8, x_17, x_64, x_10, x_11, x_12, x_13, x_14, x_15, x_63); +return x_65; } else { -lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; -x_65 = lean_ctor_get(x_59, 1); -lean_inc(x_65); -lean_dec(x_59); +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; +x_66 = lean_ctor_get(x_60, 1); +lean_inc(x_66); +lean_dec(x_60); lean_inc(x_2); -x_66 = l_Lean_MessageData_ofName(x_2); -x_67 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__8; -x_68 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_68, 0, x_67); -lean_ctor_set(x_68, 1, x_66); -x_69 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__10; -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_MessageData_ofExpr(x_52); -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_67 = l_Lean_MessageData_ofName(x_2); +x_68 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__8; +x_69 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_67); +x_70 = l_Lean_Elab_Command_elabAxiom___lambda__5___closed__10; +x_71 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_71, 0, x_69); +lean_ctor_set(x_71, 1, x_70); +x_72 = l_Lean_MessageData_ofExpr(x_53); x_73 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_67); -x_74 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_58, x_73, x_10, x_11, x_12, x_13, x_14, x_15, x_65); -x_75 = lean_ctor_get(x_74, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_74, 1); +lean_ctor_set(x_73, 0, x_71); +lean_ctor_set(x_73, 1, x_72); +x_74 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_68); +x_75 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_59, x_74, x_10, x_11, x_12, x_13, x_14, x_15, x_66); +x_76 = lean_ctor_get(x_75, 0); lean_inc(x_76); -lean_dec(x_74); -x_77 = l_Lean_Elab_Command_elabAxiom___lambda__4(x_57, x_2, x_8, x_17, x_75, x_10, x_11, x_12, x_13, x_14, x_15, x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); lean_dec(x_75); -return x_77; +x_78 = l_Lean_Elab_Command_elabAxiom___lambda__4(x_58, x_2, x_8, x_17, x_76, x_10, x_11, x_12, x_13, x_14, x_15, x_77); +lean_dec(x_76); +return x_78; } } } else { -uint8_t x_78; +uint8_t x_79; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -2505,29 +2506,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_78 = !lean_is_exclusive(x_35); -if (x_78 == 0) +x_79 = !lean_is_exclusive(x_36); +if (x_79 == 0) { -return x_35; +return x_36; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_35, 0); -x_80 = lean_ctor_get(x_35, 1); +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_36, 0); +x_81 = lean_ctor_get(x_36, 1); +lean_inc(x_81); lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_35); -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_dec(x_36); +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; } } } else { -uint8_t x_82; +uint8_t x_83; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -2542,29 +2543,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_82 = !lean_is_exclusive(x_32); -if (x_82 == 0) +x_83 = !lean_is_exclusive(x_33); +if (x_83 == 0) { -return x_32; +return x_33; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_32, 0); -x_84 = lean_ctor_get(x_32, 1); +lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_84 = lean_ctor_get(x_33, 0); +x_85 = lean_ctor_get(x_33, 1); +lean_inc(x_85); lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_32); -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_dec(x_33); +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_86; +uint8_t x_87; lean_dec(x_22); lean_dec(x_17); lean_dec(x_15); @@ -2581,29 +2582,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_86 = !lean_is_exclusive(x_25); -if (x_86 == 0) +x_87 = !lean_is_exclusive(x_26); +if (x_87 == 0) { -return x_25; +return x_26; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_25, 0); -x_88 = lean_ctor_get(x_25, 1); +lean_object* x_88; lean_object* x_89; lean_object* x_90; +x_88 = lean_ctor_get(x_26, 0); +x_89 = lean_ctor_get(x_26, 1); +lean_inc(x_89); lean_inc(x_88); -lean_inc(x_87); -lean_dec(x_25); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_dec(x_26); +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 { -uint8_t x_90; +uint8_t x_91; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -2619,29 +2620,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_90 = !lean_is_exclusive(x_21); -if (x_90 == 0) +x_91 = !lean_is_exclusive(x_21); +if (x_91 == 0) { return x_21; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_21, 0); -x_92 = lean_ctor_get(x_21, 1); +lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_92 = lean_ctor_get(x_21, 0); +x_93 = lean_ctor_get(x_21, 1); +lean_inc(x_93); lean_inc(x_92); -lean_inc(x_91); lean_dec(x_21); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_92); +lean_ctor_set(x_94, 1, x_93); +return x_94; } } } else { -uint8_t x_94; +uint8_t x_95; lean_dec(x_17); lean_dec(x_15); lean_dec(x_14); @@ -2658,23 +2659,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_94 = !lean_is_exclusive(x_19); -if (x_94 == 0) +x_95 = !lean_is_exclusive(x_19); +if (x_95 == 0) { return x_19; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_19, 0); -x_96 = lean_ctor_get(x_19, 1); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_19, 0); +x_97 = lean_ctor_get(x_19, 1); +lean_inc(x_97); lean_inc(x_96); -lean_inc(x_95); lean_dec(x_19); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +x_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; } } } diff --git a/stage0/stdlib/Lean/Elab/Deriving/BEq.c b/stage0/stdlib/Lean/Elab/Deriving/BEq.c index e9860d24c20e..69c756a9ee9d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/BEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/BEq.c @@ -32,6 +32,7 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkMutualBlock___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMutualBlock___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__6___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_BEq_0__Lean_Elab_Deriving_BEq_mkBEqEnumFun___closed__19; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__5___lambda__1___closed__4; lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -112,7 +113,6 @@ static lean_object* l_Lean_Elab_Deriving_BEq_mkAuxFunction___lambda__1___closed_ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_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*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_BEq_mkMatch_mkAlts___spec__4___lambda__1___closed__2; @@ -847,7 +847,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_BEq_mkM lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Basic.c b/stage0/stdlib/Lean/Elab/Deriving/Basic.c index a8ec6c4a5630..f27b5c4f3f1d 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Basic.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Basic.c @@ -23,6 +23,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_applyDerivingHandlers___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_processDefDeriving___spec__3___closed__1; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_2499____closed__10; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_getOptDerivingClasses___closed__2; static lean_object* l_Lean_Elab_elabDeriving___closed__8; lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -95,7 +96,6 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea static lean_object* l_Lean_Elab_elabDeriving___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_elabDeriving___lambda__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Deriving_Basic___hyg_2499____closed__16; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_elabDeriving___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Term_processDefDeriving___spec__4___rarg(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); @@ -1319,7 +1319,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_processDefDeriving(lean_object* x_1, l lean_object* x_10; lean_object* x_11; lean_object* x_22; lean_inc(x_3); lean_inc(x_2); -x_22 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_22 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_22) == 0) { lean_object* x_23; diff --git a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c index b2fed6b81375..a0314d555ef5 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/DecEq.c +++ b/stage0/stdlib/Lean/Elab/Deriving/DecEq.c @@ -39,6 +39,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqInstanceHandler___box static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__38; static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__3___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_initFn____x40_Lean_Elab_Deriving_DecEq___hyg_4708____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -163,7 +164,6 @@ lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__34; lean_object* lean_nat_div(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_DecEq_mkMatch_mkSameCtorRhs___lambda__1___closed__18; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_mkMatch_mkAlts___spec__1___closed__4; static lean_object* l_Lean_Elab_Deriving_DecEq_mkDecEqEnum___closed__56; @@ -2287,7 +2287,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_DecEq_m lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c index d951fb2b2eee..bbbf35db3e36 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c +++ b/stage0/stdlib/Lean/Elab/Deriving/FromToJson.c @@ -41,6 +41,7 @@ static lean_object* l_Array_foldrMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson static lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___lambda__4___closed__4; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_FromToJson_mkToJsonInstance_mkAlts___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda__2___closed__9; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___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_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___spec__6___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkToJsonInstance___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_Deriving_FromToJson_mkToJsonInstance___lambda__3___closed__32; @@ -179,7 +180,6 @@ static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda_ 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_Array_mapIdxM_map___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance_mkAlts___spec__3___closed__11; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_FromToJson_mkFromJsonInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance___lambda__3___closed__5; static lean_object* l_Lean_Elab_Deriving_FromToJson_mkJsonField___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_FromToJson_mkFromJsonInstance_mkAlts(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -782,7 +782,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_FromToJ lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c index 78683ee2d264..b56e9cbecf1c 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Hashable.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Hashable.c @@ -21,6 +21,7 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMa static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__5; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Deriving_Hashable_mkAuxFunction___lambda__1___closed__19; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__4___closed__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*); @@ -89,7 +90,6 @@ lean_object* l_Lean_MessageData_ofSyntax(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_List_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__5___lambda__1___closed__3; static lean_object* l___private_Lean_Elab_Deriving_Hashable_0__Lean_Elab_Deriving_Hashable_mkHashableInstanceCmds___closed__1; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Hashable_mkMatch_mkAlts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); @@ -354,7 +354,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Hashabl lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c index 4f63fae0e365..c0b7360bd976 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Inhabited.c @@ -32,6 +32,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInha static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___closed__3; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmd_x3f___spec__1___closed__1; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -116,7 +117,6 @@ lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); lean_object* l_Lean_MessageData_ofSyntax(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___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstance___closed__2; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Inhabited_0__Lean_Elab_mkInhabitedInstanceUsing_mkInstanceCmdWith___closed__20; lean_object* l_Lean_Parser_Term_implicitBinder(uint8_t); @@ -5050,7 +5050,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deri lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; @@ -5206,7 +5206,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Elab_Derivi lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c index 34a99563dbb2..3581d6bb33de 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Nonempty.c @@ -20,6 +20,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Derivin static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__49; static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__31; lean_object* lean_mk_empty_array_with_capacity(lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__30; static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__40; lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -69,7 +70,6 @@ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___spec__6___boxed(lean_object*, lean_object*, lean_object*, 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___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__19; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__34; static lean_object* l___private_Lean_Elab_Deriving_Nonempty_0__Lean_Elab_mkNonemptyInstance___lambda__1___closed__9; lean_object* l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); @@ -277,7 +277,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at___private_Lean_Elab_Deri lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Ord.c b/stage0/stdlib/Lean/Elab/Deriving/Ord.c index 7c70d8ad2138..8a9b03ba030f 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Ord.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Ord.c @@ -27,6 +27,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_isInductive___at_Lean_Elab_Deriving_Ord_mkOrdInstanceHandler___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__20; static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__24; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkMatch___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_initFn____x40_Lean_Elab_Deriving_Ord___hyg_2780____closed__11; @@ -112,7 +113,6 @@ lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); lean_object* l_Lean_MessageData_ofSyntax(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_List_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__7___lambda__2___closed__17; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Ord_mkAuxFunction___lambda__1___closed__12; lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Ord_mkMatch_mkAlts___spec__5___lambda__1___closed__20; @@ -398,7 +398,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Ord_mkM lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Deriving/Repr.c b/stage0/stdlib/Lean/Elab/Deriving/Repr.c index 0eca891f4d06..a48a767a86a4 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Repr.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Repr.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkMutualBlock___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__1; @@ -31,12 +32,15 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__14; LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___spec__1(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__6; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__2; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__3; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__32; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__2; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__17; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__9; @@ -45,7 +49,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkReprInstanceHandler(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__30; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4; lean_object* l_Lean_Name_toString(lean_object*, uint8_t); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__6; lean_object* l_Lean_MessageData_ofList(lean_object*); @@ -54,7 +57,6 @@ lean_object* lean_array_push(lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__37; static lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock___closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__22; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__14; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__6; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__9; @@ -76,13 +78,13 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkB static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__2; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__43; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__8; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1; lean_object* l_Lean_Expr_fvarId_x21(lean_object*); lean_object* lean_environment_find(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__16; static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1___closed__1; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__24; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524_(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__22; lean_object* l_Lean_Elab_Deriving_mkLocalInstanceLetDecls(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__2; @@ -91,9 +93,9 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo lean_object* l_Lean_Syntax_node5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__29; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9; lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__3; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__20; LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__33; @@ -105,6 +107,7 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__4; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__4; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__8; @@ -116,12 +119,14 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__11; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__20; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__17; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__19; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__31; size_t lean_usize_of_nat(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__5; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__25; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6; lean_object* l_Lean_Elab_registerDerivingHandler(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__4(size_t, size_t, lean_object*); static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__5; @@ -141,7 +146,6 @@ lean_object* l_Lean_MessageData_ofSyntax(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* 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_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__14; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__24; 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*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,9 +165,9 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___clos lean_object* l_Lean_FVarId_getDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__17; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__32; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__29; lean_object* l_Lean_Elab_Command_liftTermElabM___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkMutualBlock(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -186,11 +190,11 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu lean_object* l_Lean_Elab_Deriving_mkHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__5; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388_(lean_object*); extern lean_object* l_Lean_instInhabitedExpr; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForInduct___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___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_throwError___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__7; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__10; @@ -198,7 +202,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__20; lean_object* l_Lean_Elab_Deriving_mkContext(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__27; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__1___closed__1; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); @@ -239,7 +242,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_mkBody(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__10; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__1; @@ -260,6 +262,7 @@ static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Re 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_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2; 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*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__1; @@ -283,7 +286,7 @@ uint8_t l_Lean_isStructure(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__1; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__27; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__39; uint8_t l_Lean_LocalDecl_binderInfo(lean_object*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__2; @@ -292,32 +295,31 @@ static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForIndu lean_object* l_List_reverse___rarg(lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__16; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__15; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__16; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__18; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__3; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__12; size_t lean_usize_add(size_t, size_t); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__11; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedName; static lean_object* l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__2; lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__30; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__26; +static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16; extern lean_object* l_Lean_instInhabitedInductiveVal; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__36; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__1; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__12; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__38; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__2; static lean_object* l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__7; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__28; lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__1___closed__2; @@ -333,7 +335,6 @@ static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__13; static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__34; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__5; -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12; static lean_object* l_Lean_Elab_Deriving_Repr_mkAuxFunction___lambda__1___closed__4; lean_object* l_String_toSubstring_x27(lean_object*); static lean_object* l_List_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__3___lambda__1___closed__5; @@ -353,7 +354,6 @@ static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyFo static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__1___closed__6; static lean_object* l_Lean_Elab_Deriving_Repr_mkBodyForStruct___lambda__2___closed__3; lean_object* l_Lean_Syntax_mkNumLit(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Deriving_Repr_mkReprInstanceHandler___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1() { _start: @@ -904,7 +904,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at_Lean_Elab_Deriving_Repr_mk lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; @@ -3145,7 +3145,7 @@ x_23 = lean_unsigned_to_nat(0u); x_24 = lean_nat_dec_eq(x_10, 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_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; lean_object* x_135; +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_181; lean_object* x_182; uint8_t x_183; lean_object* x_184; lean_object* x_185; x_25 = lean_unsigned_to_nat(1u); x_26 = lean_nat_sub(x_10, x_25); lean_dec(x_10); @@ -3161,42 +3161,42 @@ if (lean_is_exclusive(x_14)) { lean_dec_ref(x_14); x_35 = lean_box(0); } -x_131 = lean_nat_add(x_9, x_11); -x_132 = lean_array_get_size(x_3); -x_133 = lean_nat_dec_lt(x_131, x_132); -lean_dec(x_132); -x_134 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__9; -x_135 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_134, x_19, x_20, x_21); -if (x_133 == 0) -{ -lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; -lean_dec(x_131); -x_136 = lean_ctor_get(x_135, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_135, 1); -lean_inc(x_137); -lean_dec(x_135); -x_138 = l_Lean_instInhabitedExpr; -x_139 = l___private_Init_GetElem_0__outOfBounds___rarg(x_138); -x_36 = x_139; -x_37 = x_136; -x_38 = x_137; -goto block_130; +x_181 = lean_nat_add(x_9, x_11); +x_182 = lean_array_get_size(x_3); +x_183 = lean_nat_dec_lt(x_181, x_182); +lean_dec(x_182); +x_184 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__9; +x_185 = l___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(x_184, x_19, x_20, x_21); +if (x_183 == 0) +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_181); +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +lean_dec(x_185); +x_188 = l_Lean_instInhabitedExpr; +x_189 = l___private_Init_GetElem_0__outOfBounds___rarg(x_188); +x_36 = x_189; +x_37 = x_186; +x_38 = x_187; +goto block_180; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_135, 0); -lean_inc(x_140); -x_141 = lean_ctor_get(x_135, 1); -lean_inc(x_141); -lean_dec(x_135); -x_142 = lean_array_fget(x_3, x_131); -lean_dec(x_131); -x_36 = x_142; -x_37 = x_140; -x_38 = x_141; -goto block_130; +lean_object* x_190; lean_object* x_191; lean_object* x_192; +x_190 = lean_ctor_get(x_185, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_185, 1); +lean_inc(x_191); +lean_dec(x_185); +x_192 = lean_array_fget(x_3, x_181); +lean_dec(x_181); +x_36 = x_192; +x_37 = x_190; +x_38 = x_191; +goto block_180; } block_32: { @@ -3212,244 +3212,280 @@ x_14 = x_29; x_21 = x_28; goto _start; } -block_130: +block_180: { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; lean_object* x_108; lean_object* x_109; x_39 = lean_mk_syntax_ident(x_37); lean_inc(x_39); x_40 = lean_array_push(x_33, x_39); lean_inc(x_36); -x_41 = l_Lean_Expr_fvarId_x21(x_36); +x_108 = l_Lean_Expr_fvarId_x21(x_36); lean_inc(x_17); -x_42 = l_Lean_FVarId_getDecl(x_41, x_17, x_18, x_19, x_20, x_38); -if (lean_obj_tag(x_42) == 0) +x_109 = l_Lean_FVarId_getDecl(x_108, x_17, x_18, x_19, x_20, x_38); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_43; lean_object* x_44; uint8_t x_45; uint8_t x_46; -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_LocalDecl_binderInfo(x_43); -lean_dec(x_43); -x_46 = l_Lean_BinderInfo_isExplicit(x_45); -if (x_46 == 0) +lean_object* x_110; lean_object* x_111; uint8_t x_112; uint8_t x_113; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = l_Lean_LocalDecl_binderInfo(x_110); +lean_dec(x_110); +x_113 = l_Lean_BinderInfo_isExplicit(x_112); +if (x_113 == 0) { -lean_object* x_47; lean_object* x_48; +lean_object* x_114; lean_object* x_115; lean_dec(x_39); lean_dec(x_36); -if (lean_is_scalar(x_35)) { - x_47 = lean_alloc_ctor(0, 2, 0); -} else { - x_47 = x_35; -} -lean_ctor_set(x_47, 0, x_40); -lean_ctor_set(x_47, 1, x_34); -x_48 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_48, 0, x_47); -x_27 = x_48; -x_28 = x_44; +lean_dec(x_35); +x_114 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_114, 0, x_40); +lean_ctor_set(x_114, 1, x_34); +x_115 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_115, 0, x_114); +x_27 = x_115; +x_28 = x_111; goto block_32; } else { -lean_object* x_49; +lean_object* x_116; lean_inc(x_20); lean_inc(x_19); lean_inc(x_18); lean_inc(x_17); -x_49 = lean_infer_type(x_36, x_17, x_18, x_19, x_20, x_44); -if (lean_obj_tag(x_49) == 0) +lean_inc(x_36); +x_116 = lean_infer_type(x_36, x_17, x_18, x_19, x_20, x_111); +if (lean_obj_tag(x_116) == 0) +{ +lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; uint8_t x_121; +x_117 = lean_ctor_get(x_116, 0); +lean_inc(x_117); +x_118 = lean_ctor_get(x_116, 1); +lean_inc(x_118); +lean_dec(x_116); +x_119 = lean_ctor_get(x_1, 0); +x_120 = lean_ctor_get(x_119, 0); +x_121 = l_Lean_Expr_isAppOf(x_117, x_120); +lean_dec(x_117); +if (x_121 == 0) +{ +lean_object* x_122; +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_36); +x_122 = l_Lean_Meta_isType(x_36, x_17, x_18, x_19, x_20, x_118); +if (lean_obj_tag(x_122) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; uint8_t x_54; -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_1, 0); -x_53 = lean_ctor_get(x_52, 0); -x_54 = l_Lean_Expr_isAppOf(x_50, x_53); -lean_dec(x_50); -if (x_54 == 0) +lean_object* x_123; uint8_t x_124; +x_123 = lean_ctor_get(x_122, 0); +lean_inc(x_123); +x_124 = lean_unbox(x_123); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; +lean_dec(x_123); +x_125 = lean_ctor_get(x_122, 1); +lean_inc(x_125); +lean_dec(x_122); +lean_inc(x_20); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +x_126 = l_Lean_Meta_isProof(x_36, x_17, x_18, x_19, x_20, x_125); +if (lean_obj_tag(x_126) == 0) { -lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_55 = lean_ctor_get(x_19, 5); -lean_inc(x_55); -x_56 = 0; -x_57 = l_Lean_SourceInfo_fromRef(x_55, x_56); -x_58 = lean_ctor_get(x_19, 10); -lean_inc(x_58); -x_59 = lean_st_ref_get(x_20, x_51); -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_ctor_get(x_60, 0); -lean_inc(x_62); -lean_dec(x_60); -x_63 = lean_environment_main_module(x_62); -x_64 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; -lean_inc(x_57); -x_65 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_65, 0, x_57); -lean_ctor_set(x_65, 1, x_64); -x_66 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; -lean_inc(x_58); -lean_inc(x_63); -x_67 = l_Lean_addMacroScope(x_63, x_66, x_58); -x_68 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__6; -lean_inc(x_6); -x_69 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_69, 0, x_68); -lean_ctor_set(x_69, 1, x_6); -x_70 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__7; -lean_inc(x_7); -x_71 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_71, 0, x_70); -lean_ctor_set(x_71, 1, x_7); -x_72 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_72, 0, x_69); -lean_ctor_set(x_72, 1, x_71); -x_73 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_57); -x_74 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_74, 0, x_57); -lean_ctor_set(x_74, 1, x_73); -lean_ctor_set(x_74, 2, x_67); -lean_ctor_set(x_74, 3, x_72); -x_75 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -lean_inc(x_65); -lean_inc(x_57); -x_76 = l_Lean_Syntax_node3(x_57, x_75, x_34, x_65, x_74); -x_77 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; -x_78 = l_Lean_addMacroScope(x_63, x_77, x_58); -lean_inc(x_6); -x_79 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_6); -x_80 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; -lean_inc(x_7); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_80); -lean_ctor_set(x_81, 1, x_7); -x_82 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_82, 0, x_79); -lean_ctor_set(x_82, 1, x_81); -x_83 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; -lean_inc(x_57); -x_84 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_84, 0, x_57); -lean_ctor_set(x_84, 1, x_83); -lean_ctor_set(x_84, 2, x_78); -lean_ctor_set(x_84, 3, x_82); -lean_inc(x_8); -lean_inc(x_57); -x_85 = l_Lean_Syntax_node1(x_57, x_8, x_39); -lean_inc(x_5); -lean_inc(x_57); -x_86 = l_Lean_Syntax_node2(x_57, x_5, x_84, x_85); -x_87 = l_Lean_Syntax_node3(x_57, x_75, x_76, x_65, x_86); -if (lean_is_scalar(x_35)) { - x_88 = lean_alloc_ctor(0, 2, 0); -} else { - x_88 = x_35; +lean_object* x_127; lean_object* x_128; uint8_t x_129; +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = lean_unbox(x_127); +lean_dec(x_127); +x_41 = x_129; +x_42 = x_128; +goto block_107; } -lean_ctor_set(x_88, 0, x_40); -lean_ctor_set(x_88, 1, x_87); -x_89 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_89, 0, x_88); -x_27 = x_89; -x_28 = x_61; -goto block_32; +else +{ +uint8_t x_130; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_130 = !lean_is_exclusive(x_126); +if (x_130 == 0) +{ +return x_126; } else { -lean_object* x_90; uint8_t x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_90 = lean_ctor_get(x_19, 5); -lean_inc(x_90); -x_91 = 0; -x_92 = l_Lean_SourceInfo_fromRef(x_90, x_91); -x_93 = lean_ctor_get(x_19, 10); -lean_inc(x_93); -x_94 = lean_st_ref_get(x_20, x_51); -x_95 = lean_ctor_get(x_94, 0); -lean_inc(x_95); -x_96 = lean_ctor_get(x_94, 1); -lean_inc(x_96); -lean_dec(x_94); -x_97 = lean_ctor_get(x_95, 0); -lean_inc(x_97); -lean_dec(x_95); -x_98 = lean_environment_main_module(x_97); -x_99 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; -lean_inc(x_92); -x_100 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_100, 0, x_92); -lean_ctor_set(x_100, 1, x_99); -x_101 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; -x_102 = l_Lean_addMacroScope(x_98, x_101, x_93); -x_103 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__6; +lean_object* x_131; lean_object* x_132; lean_object* x_133; +x_131 = lean_ctor_get(x_126, 0); +x_132 = lean_ctor_get(x_126, 1); +lean_inc(x_132); +lean_inc(x_131); +lean_dec(x_126); +x_133 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_133, 0, x_131); +lean_ctor_set(x_133, 1, x_132); +return x_133; +} +} +} +else +{ +lean_object* x_134; uint8_t x_135; +lean_dec(x_36); +x_134 = lean_ctor_get(x_122, 1); +lean_inc(x_134); +lean_dec(x_122); +x_135 = lean_unbox(x_123); +lean_dec(x_123); +x_41 = x_135; +x_42 = x_134; +goto block_107; +} +} +else +{ +uint8_t x_136; +lean_dec(x_40); +lean_dec(x_39); +lean_dec(x_36); +lean_dec(x_35); +lean_dec(x_34); +lean_dec(x_26); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_11); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_2); +x_136 = !lean_is_exclusive(x_122); +if (x_136 == 0) +{ +return x_122; +} +else +{ +lean_object* x_137; lean_object* x_138; lean_object* x_139; +x_137 = lean_ctor_get(x_122, 0); +x_138 = lean_ctor_get(x_122, 1); +lean_inc(x_138); +lean_inc(x_137); +lean_dec(x_122); +x_139 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_139, 0, x_137); +lean_ctor_set(x_139, 1, x_138); +return x_139; +} +} +} +else +{ +lean_object* x_140; uint8_t x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; +lean_dec(x_36); +lean_dec(x_35); +x_140 = lean_ctor_get(x_19, 5); +lean_inc(x_140); +x_141 = 0; +x_142 = l_Lean_SourceInfo_fromRef(x_140, x_141); +x_143 = lean_ctor_get(x_19, 10); +lean_inc(x_143); +x_144 = lean_st_ref_get(x_20, x_118); +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +lean_dec(x_144); +x_147 = lean_ctor_get(x_145, 0); +lean_inc(x_147); +lean_dec(x_145); +x_148 = lean_environment_main_module(x_147); +x_149 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; +lean_inc(x_142); +x_150 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_150, 0, x_142); +lean_ctor_set(x_150, 1, x_149); +x_151 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; +x_152 = l_Lean_addMacroScope(x_148, x_151, x_143); +x_153 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__6; lean_inc(x_6); -x_104 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_104, 0, x_103); -lean_ctor_set(x_104, 1, x_6); -x_105 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__7; +x_154 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_154, 0, x_153); +lean_ctor_set(x_154, 1, x_6); +x_155 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__7; lean_inc(x_7); -x_106 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_106, 0, x_105); -lean_ctor_set(x_106, 1, x_7); -x_107 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_107, 0, x_104); -lean_ctor_set(x_107, 1, x_106); -x_108 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; -lean_inc(x_92); -x_109 = lean_alloc_ctor(3, 4, 0); -lean_ctor_set(x_109, 0, x_92); -lean_ctor_set(x_109, 1, x_108); -lean_ctor_set(x_109, 2, x_102); -lean_ctor_set(x_109, 3, x_107); -x_110 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; -lean_inc(x_100); -lean_inc(x_92); -x_111 = l_Lean_Syntax_node3(x_92, x_110, x_34, x_100, x_109); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_155); +lean_ctor_set(x_156, 1, x_7); +x_157 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_157, 0, x_154); +lean_ctor_set(x_157, 1, x_156); +x_158 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; +lean_inc(x_142); +x_159 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_159, 0, x_142); +lean_ctor_set(x_159, 1, x_158); +lean_ctor_set(x_159, 2, x_152); +lean_ctor_set(x_159, 3, x_157); +x_160 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +lean_inc(x_150); +lean_inc(x_142); +x_161 = l_Lean_Syntax_node3(x_142, x_160, x_34, x_150, x_159); lean_inc(x_2); -x_112 = lean_mk_syntax_ident(x_2); -x_113 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; -lean_inc(x_92); -x_114 = lean_alloc_ctor(2, 2, 0); -lean_ctor_set(x_114, 0, x_92); -lean_ctor_set(x_114, 1, x_113); -x_115 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__6; -lean_inc(x_92); -x_116 = l_Lean_Syntax_node1(x_92, x_115, x_114); +x_162 = lean_mk_syntax_ident(x_2); +x_163 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__7; +lean_inc(x_142); +x_164 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_164, 0, x_142); +lean_ctor_set(x_164, 1, x_163); +x_165 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__6; +lean_inc(x_142); +x_166 = l_Lean_Syntax_node1(x_142, x_165, x_164); lean_inc(x_8); -lean_inc(x_92); -x_117 = l_Lean_Syntax_node2(x_92, x_8, x_39, x_116); +lean_inc(x_142); +x_167 = l_Lean_Syntax_node2(x_142, x_8, x_39, x_166); lean_inc(x_5); -lean_inc(x_92); -x_118 = l_Lean_Syntax_node2(x_92, x_5, x_112, x_117); -x_119 = l_Lean_Syntax_node3(x_92, x_110, x_111, x_100, x_118); -if (lean_is_scalar(x_35)) { - x_120 = lean_alloc_ctor(0, 2, 0); -} else { - x_120 = x_35; -} -lean_ctor_set(x_120, 0, x_40); -lean_ctor_set(x_120, 1, x_119); -x_121 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_121, 0, x_120); -x_27 = x_121; -x_28 = x_96; +lean_inc(x_142); +x_168 = l_Lean_Syntax_node2(x_142, x_5, x_162, x_167); +x_169 = l_Lean_Syntax_node3(x_142, x_160, x_161, x_150, x_168); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_40); +lean_ctor_set(x_170, 1, x_169); +x_171 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_171, 0, x_170); +x_27 = x_171; +x_28 = x_146; goto block_32; } } else { -uint8_t x_122; +uint8_t x_172; lean_dec(x_40); lean_dec(x_39); +lean_dec(x_36); lean_dec(x_35); lean_dec(x_34); lean_dec(x_26); @@ -3463,30 +3499,30 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_122 = !lean_is_exclusive(x_49); -if (x_122 == 0) +x_172 = !lean_is_exclusive(x_116); +if (x_172 == 0) { -return x_49; +return x_116; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_49, 0); -x_124 = lean_ctor_get(x_49, 1); -lean_inc(x_124); -lean_inc(x_123); -lean_dec(x_49); -x_125 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_125, 0, x_123); -lean_ctor_set(x_125, 1, x_124); -return x_125; +lean_object* x_173; lean_object* x_174; lean_object* x_175; +x_173 = lean_ctor_get(x_116, 0); +x_174 = lean_ctor_get(x_116, 1); +lean_inc(x_174); +lean_inc(x_173); +lean_dec(x_116); +x_175 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_175, 0, x_173); +lean_ctor_set(x_175, 1, x_174); +return x_175; } } } } else { -uint8_t x_126; +uint8_t x_176; lean_dec(x_40); lean_dec(x_39); lean_dec(x_36); @@ -3503,30 +3539,199 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_126 = !lean_is_exclusive(x_42); -if (x_126 == 0) +x_176 = !lean_is_exclusive(x_109); +if (x_176 == 0) { -return x_42; +return x_109; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_42, 0); -x_128 = lean_ctor_get(x_42, 1); -lean_inc(x_128); -lean_inc(x_127); -lean_dec(x_42); -x_129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -return x_129; +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_109, 0); +x_178 = lean_ctor_get(x_109, 1); +lean_inc(x_178); +lean_inc(x_177); +lean_dec(x_109); +x_179 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_179, 0, x_177); +lean_ctor_set(x_179, 1, x_178); +return x_179; +} +} +block_107: +{ +if (x_41 == 0) +{ +lean_object* x_43; uint8_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; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_43 = lean_ctor_get(x_19, 5); +lean_inc(x_43); +x_44 = 0; +x_45 = l_Lean_SourceInfo_fromRef(x_43, x_44); +x_46 = lean_ctor_get(x_19, 10); +lean_inc(x_46); +x_47 = lean_st_ref_get(x_20, x_42); +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); +lean_dec(x_47); +x_50 = lean_ctor_get(x_48, 0); +lean_inc(x_50); +lean_dec(x_48); +x_51 = lean_environment_main_module(x_50); +x_52 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; +lean_inc(x_45); +x_53 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_53, 0, x_45); +lean_ctor_set(x_53, 1, x_52); +x_54 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; +lean_inc(x_46); +lean_inc(x_51); +x_55 = l_Lean_addMacroScope(x_51, x_54, x_46); +x_56 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__6; +lean_inc(x_6); +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_6); +x_58 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__7; +lean_inc(x_7); +x_59 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_59, 0, x_58); +lean_ctor_set(x_59, 1, x_7); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_59); +x_61 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; +lean_inc(x_45); +x_62 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_62, 0, x_45); +lean_ctor_set(x_62, 1, x_61); +lean_ctor_set(x_62, 2, x_55); +lean_ctor_set(x_62, 3, x_60); +x_63 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +lean_inc(x_53); +lean_inc(x_45); +x_64 = l_Lean_Syntax_node3(x_45, x_63, x_34, x_53, x_62); +x_65 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__3; +x_66 = l_Lean_addMacroScope(x_51, x_65, x_46); +lean_inc(x_6); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_65); +lean_ctor_set(x_67, 1, x_6); +x_68 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__4; +lean_inc(x_7); +x_69 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_69, 0, x_68); +lean_ctor_set(x_69, 1, x_7); +x_70 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_70, 0, x_67); +lean_ctor_set(x_70, 1, x_69); +x_71 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForInduct_mkAlts___spec__2___closed__2; +lean_inc(x_45); +x_72 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_72, 0, x_45); +lean_ctor_set(x_72, 1, x_71); +lean_ctor_set(x_72, 2, x_66); +lean_ctor_set(x_72, 3, x_70); +lean_inc(x_8); +lean_inc(x_45); +x_73 = l_Lean_Syntax_node1(x_45, x_8, x_39); +lean_inc(x_5); +lean_inc(x_45); +x_74 = l_Lean_Syntax_node2(x_45, x_5, x_72, x_73); +x_75 = l_Lean_Syntax_node3(x_45, x_63, x_64, x_53, x_74); +if (lean_is_scalar(x_35)) { + x_76 = lean_alloc_ctor(0, 2, 0); +} else { + x_76 = x_35; +} +lean_ctor_set(x_76, 0, x_40); +lean_ctor_set(x_76, 1, x_75); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_76); +x_27 = x_77; +x_28 = x_49; +goto block_32; +} +else +{ +lean_object* x_78; uint8_t x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +lean_dec(x_39); +x_78 = lean_ctor_get(x_19, 5); +lean_inc(x_78); +x_79 = 0; +x_80 = l_Lean_SourceInfo_fromRef(x_78, x_79); +x_81 = lean_ctor_get(x_19, 10); +lean_inc(x_81); +x_82 = lean_st_ref_get(x_20, x_42); +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); +lean_dec(x_83); +x_86 = lean_environment_main_module(x_85); +x_87 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__5; +lean_inc(x_80); +x_88 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 1, x_87); +x_89 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__5; +x_90 = l_Lean_addMacroScope(x_86, x_89, x_81); +x_91 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__6; +lean_inc(x_6); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_91); +lean_ctor_set(x_92, 1, x_6); +x_93 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__7; +lean_inc(x_7); +x_94 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_94, 0, x_93); +lean_ctor_set(x_94, 1, x_7); +x_95 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_94); +x_96 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___closed__3; +lean_inc(x_80); +x_97 = lean_alloc_ctor(3, 4, 0); +lean_ctor_set(x_97, 0, x_80); +lean_ctor_set(x_97, 1, x_96); +lean_ctor_set(x_97, 2, x_90); +lean_ctor_set(x_97, 3, x_95); +x_98 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__4; +lean_inc(x_88); +lean_inc(x_80); +x_99 = l_Lean_Syntax_node3(x_80, x_98, x_34, x_88, x_97); +x_100 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__34; +lean_inc(x_80); +x_101 = lean_alloc_ctor(2, 2, 0); +lean_ctor_set(x_101, 0, x_80); +lean_ctor_set(x_101, 1, x_100); +x_102 = l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_Repr_mkBodyForStruct___spec__3___lambda__1___closed__7; +lean_inc(x_80); +x_103 = l_Lean_Syntax_node1(x_80, x_102, x_101); +x_104 = l_Lean_Syntax_node3(x_80, x_98, x_99, x_88, x_103); +if (lean_is_scalar(x_35)) { + x_105 = lean_alloc_ctor(0, 2, 0); +} else { + x_105 = x_35; +} +lean_ctor_set(x_105, 0, x_40); +lean_ctor_set(x_105, 1, x_104); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_105); +x_27 = x_106; +x_28 = x_84; +goto block_32; } } } } else { -lean_object* x_143; +lean_object* x_193; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -3538,15 +3743,15 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_143 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_143, 0, x_14); -lean_ctor_set(x_143, 1, x_21); -return x_143; +x_193 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_193, 0, x_14); +lean_ctor_set(x_193, 1, x_21); +return x_193; } } else { -lean_object* x_144; +lean_object* x_194; lean_dec(x_20); lean_dec(x_19); lean_dec(x_18); @@ -3558,10 +3763,10 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_14); -lean_ctor_set(x_144, 1, x_21); -return x_144; +x_194 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_194, 0, x_14); +lean_ctor_set(x_194, 1, x_21); +return x_194; } } } @@ -7260,7 +7465,7 @@ lean_dec(x_1); return x_5; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1() { _start: { lean_object* x_1; @@ -7268,7 +7473,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Deriving_Repr_mkReprInstanceHandler return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -7278,37 +7483,37 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2; x_2 = l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3; x_2 = l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6() { _start: { lean_object* x_1; @@ -7316,17 +7521,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5; -x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5; +x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8() { _start: { lean_object* x_1; @@ -7334,57 +7539,57 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7; -x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7; +x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10; x_2 = l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11; x_2 = l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14() { _start: { lean_object* x_1; @@ -7392,32 +7597,32 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13; -x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14; +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13; +x_2 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16() { +static lean_object* _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15; -x_2 = lean_unsigned_to_nat(3388u); +x_1 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15; +x_2 = lean_unsigned_to_nat(3524u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; x_2 = l_Lean_Elab_Deriving_Repr_mkReprHeader___closed__2; -x_3 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1; +x_3 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1; x_4 = l_Lean_Elab_registerDerivingHandler(x_2, x_3, x_1); if (lean_obj_tag(x_4) == 0) { @@ -7427,7 +7632,7 @@ lean_inc(x_5); lean_dec(x_4); x_6 = l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__4; x_7 = 0; -x_8 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16; +x_8 = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16; x_9 = l_Lean_registerTraceClass(x_6, x_7, x_8, x_5); return x_9; } @@ -7864,39 +8069,39 @@ l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd lean_mark_persistent(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__7); l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__8 = _init_l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__8(); lean_mark_persistent(l___private_Lean_Elab_Deriving_Repr_0__Lean_Elab_Deriving_Repr_mkReprInstanceCmd___closed__8); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__1); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__2); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__3); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__4); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__5); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__6); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__7); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__8); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__9); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__10); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__11); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__12); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__13); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__14); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__15); -l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16(); -lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388____closed__16); -if (builtin) {res = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3388_(lean_io_mk_world()); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__1); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__2); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__3); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__4); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__5); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__6); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__7); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__8); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__9); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__10); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__11); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__12); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__13); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__14); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__15); +l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16 = _init_l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16(); +lean_mark_persistent(l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524____closed__16); +if (builtin) {res = l_Lean_Elab_Deriving_Repr_initFn____x40_Lean_Elab_Deriving_Repr___hyg_3524_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Deriving/Util.c b/stage0/stdlib/Lean/Elab/Deriving/Util.c index 6e4075271ccc..13bcd3472c73 100644 --- a/stage0/stdlib/Lean/Elab/Deriving/Util.c +++ b/stage0/stdlib/Lean/Elab/Deriving/Util.c @@ -25,6 +25,7 @@ lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_mkContext___spec__1___closed__3; static lean_object* l_Lean_Elab_Deriving_mkInductiveApp___closed__1; static lean_object* l_Lean_Elab_Deriving_mkInductiveApp___closed__2; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Deriving_mkContext___closed__4; static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkInstanceCmds___spec__1___lambda__1___closed__9; static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Deriving_mkImplicitBinders___spec__1___closed__2; @@ -84,7 +85,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Deriving_mkLet(lean_object*, lean_object*, 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*); static lean_object* l_Lean_Elab_Deriving_mkInductiveApp___closed__11; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_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*); lean_object* l_Lean_Parser_Term_implicitBinder(uint8_t); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Deriving_mkLocalInstanceLetDecls___spec__1___closed__12; @@ -1466,7 +1466,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoInduct___at_Lean_Elab_Deriving_mkCon lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Elab/Eval.c b/stage0/stdlib/Lean/Elab/Eval.c index fe604577996e..dda3b9885dd8 100644 --- a/stage0/stdlib/Lean/Elab/Eval.c +++ b/stage0/stdlib/Lean/Elab/Eval.c @@ -495,150 +495,151 @@ lean_inc(x_4); x_34 = l_Lean_Elab_Term_elabTermEnsuringType(x_2, x_11, x_33, x_33, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_15); if (lean_obj_tag(x_34) == 0) { -lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; +lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; 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 = 0; +x_37 = 1; +x_38 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_38 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_37, x_37, x_4, x_5, x_6, x_7, x_8, x_9, x_36); -if (lean_obj_tag(x_38) == 0) -{ -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_39 = lean_ctor_get(x_38, 1); -lean_inc(x_39); -lean_dec(x_38); -x_40 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_39); -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -x_42 = lean_ctor_get(x_40, 1); +x_39 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_37, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_36); +if (lean_obj_tag(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; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_35, x_4, x_5, x_6, x_7, x_8, x_9, x_40); +x_42 = lean_ctor_get(x_41, 0); lean_inc(x_42); -lean_dec(x_40); -lean_inc(x_41); -x_43 = l_Lean_Meta_getMVars(x_41, x_6, x_7, x_8, x_9, x_42); -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_43, 1); +x_43 = lean_ctor_get(x_41, 1); +lean_inc(x_43); +lean_dec(x_41); +lean_inc(x_42); +x_44 = l_Lean_Meta_getMVars(x_42, x_6, x_7, x_8, x_9, x_43); +x_45 = lean_ctor_get(x_44, 0); lean_inc(x_45); -lean_dec(x_43); +x_46 = lean_ctor_get(x_44, 1); +lean_inc(x_46); +lean_dec(x_44); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_46 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_44, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_45); -lean_dec(x_44); -if (lean_obj_tag(x_46) == 0) -{ -lean_object* x_47; uint8_t x_48; -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_unbox(x_47); +x_47 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_45, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_46); +lean_dec(x_45); +if (lean_obj_tag(x_47) == 0) +{ +lean_object* x_48; uint8_t x_49; +x_48 = lean_ctor_get(x_47, 0); +lean_inc(x_48); +x_49 = lean_unbox(x_48); +lean_dec(x_48); +if (x_49 == 0) +{ +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_47, 1); +lean_inc(x_50); lean_dec(x_47); -if (x_48 == 0) -{ -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_46, 1); -lean_inc(x_49); -lean_dec(x_46); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_50 = l_Lean_Meta_evalExpr___rarg(x_1, x_41, x_3, x_6, x_7, x_8, x_9, x_49); -if (lean_obj_tag(x_50) == 0) +x_51 = l_Lean_Meta_evalExpr___rarg(x_1, x_42, x_3, x_6, x_7, x_8, x_9, x_50); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_50, 0); -lean_inc(x_51); -x_52 = lean_ctor_get(x_50, 1); +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_51, 0); lean_inc(x_52); -lean_dec(x_50); -x_17 = x_51; -x_18 = x_52; +x_53 = lean_ctor_get(x_51, 1); +lean_inc(x_53); +lean_dec(x_51); +x_17 = x_52; +x_18 = x_53; goto block_24; } else { -lean_object* x_53; lean_object* x_54; -x_53 = lean_ctor_get(x_50, 0); -lean_inc(x_53); -x_54 = lean_ctor_get(x_50, 1); +lean_object* x_54; lean_object* x_55; +x_54 = lean_ctor_get(x_51, 0); lean_inc(x_54); -lean_dec(x_50); -x_25 = x_53; -x_26 = x_54; +x_55 = lean_ctor_get(x_51, 1); +lean_inc(x_55); +lean_dec(x_51); +x_25 = x_54; +x_26 = x_55; goto block_32; } } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_41); +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_dec(x_42); lean_dec(x_1); -x_55 = lean_ctor_get(x_46, 1); -lean_inc(x_55); -lean_dec(x_46); -x_56 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_55); -x_57 = lean_ctor_get(x_56, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_56, 1); +x_56 = lean_ctor_get(x_47, 1); +lean_inc(x_56); +lean_dec(x_47); +x_57 = l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(x_56); +x_58 = lean_ctor_get(x_57, 0); lean_inc(x_58); -lean_dec(x_56); -x_25 = x_57; -x_26 = x_58; +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_25 = x_58; +x_26 = x_59; goto block_32; } } else { -lean_object* x_59; lean_object* x_60; -lean_dec(x_41); +lean_object* x_60; lean_object* x_61; +lean_dec(x_42); lean_dec(x_1); -x_59 = lean_ctor_get(x_46, 0); -lean_inc(x_59); -x_60 = lean_ctor_get(x_46, 1); +x_60 = lean_ctor_get(x_47, 0); lean_inc(x_60); -lean_dec(x_46); -x_25 = x_59; -x_26 = x_60; +x_61 = lean_ctor_get(x_47, 1); +lean_inc(x_61); +lean_dec(x_47); +x_25 = x_60; +x_26 = x_61; goto block_32; } } else { -lean_object* x_61; lean_object* x_62; +lean_object* x_62; lean_object* x_63; lean_dec(x_35); lean_dec(x_1); -x_61 = lean_ctor_get(x_38, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_38, 1); +x_62 = lean_ctor_get(x_39, 0); lean_inc(x_62); -lean_dec(x_38); -x_25 = x_61; -x_26 = x_62; +x_63 = lean_ctor_get(x_39, 1); +lean_inc(x_63); +lean_dec(x_39); +x_25 = x_62; +x_26 = x_63; goto block_32; } } else { -lean_object* x_63; lean_object* x_64; +lean_object* x_64; lean_object* x_65; lean_dec(x_1); -x_63 = lean_ctor_get(x_34, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_34, 1); +x_64 = lean_ctor_get(x_34, 0); lean_inc(x_64); +x_65 = lean_ctor_get(x_34, 1); +lean_inc(x_65); lean_dec(x_34); -x_25 = x_63; -x_26 = x_64; +x_25 = x_64; +x_26 = x_65; goto block_32; } block_24: diff --git a/stage0/stdlib/Lean/Elab/Extra.c b/stage0/stdlib/Lean/Elab/Extra.c index 41d00c00d15d..6ba93163755f 100644 --- a/stage0/stdlib/Lean/Elab/Extra.c +++ b/stage0/stdlib/Lean/Elab/Extra.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn___closed__2; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___closed__3; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -24,7 +25,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkUnOp lean_object* l_Lean_Meta_mkSorry(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_KeyedDeclsAttribute_addBuiltin___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange___closed__3; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__5; @@ -34,12 +34,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Elab_Term_Op_ela 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_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp___closed__3; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_isUnknown(lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__3___closed__1; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp(lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__10; static lean_object* l_Lean_Elab_Term_elabForIn___closed__12; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__2; @@ -68,6 +68,7 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn(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*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange___closed__6; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__2; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -76,12 +77,10 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange___close static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__4; static lean_object* l_Lean_Elab_Term_elabForIn___closed__19; -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(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_Extra_0__Lean_Elab_Term_getMonadForIn___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkBinOp(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__14; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___closed__2; lean_object* l_Lean_Level_succ___override(lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); @@ -92,14 +91,17 @@ lean_object* l_Lean_Meta_trySynthInstance(lean_object*, lean_object*, lean_objec static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange___closed__6; lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); +static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___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_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___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_Extra_0__Lean_Elab_Term_Op_applyCoe_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___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange(lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_toCtorIdx___boxed(lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__9; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__5; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp___closed__4; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -109,8 +111,8 @@ LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Extra_0__Lea static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__4; static lean_object* l_Lean_Elab_Term_elabForIn___closed__7; -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2(lean_object*); uint8_t l_Lean_Elab_Term_hasCDot(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__22; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn___closed__3; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4; @@ -126,9 +128,7 @@ lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1 LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHomogeneousInstance(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*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static double l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange___closed__5; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); @@ -144,13 +144,16 @@ uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__2___closed__6; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__5; +static double l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1; lean_object* l_Lean_stringToMessageData(lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Util_Trace_0__Lean_checkTraceOption(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__25; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact___closed__1; lean_object* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy___closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__2(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -161,12 +164,12 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkBinO lean_object* l_Lean_Elab_Term_withPushMacroExpansionStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_hasMVar(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange___closed__1; +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinOp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___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_Extra_0__Lean_Elab_Term_Op_applyCoe(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion___rarg___closed__1; @@ -196,6 +199,7 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange___closed__3; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__5; +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel___closed__4; lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__4; @@ -204,15 +208,13 @@ static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0 static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__2; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__8; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__4; -lean_object* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr(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_Term_traceAtCmdPos___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4; lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel___closed__1; lean_object* l_Lean_Name_getPrefix(lean_object*); @@ -225,8 +227,8 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree static lean_object* l_Lean_Elab_Term_elabForIn___closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1(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_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__3; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__1; static lean_object* l_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__2; @@ -234,11 +236,14 @@ LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstances___at___private_Lean_Ela static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__8; static lean_object* l_Lean_Elab_Term_elabForIn___closed__4; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp(lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13; 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_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1; static lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel___closed__3; lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange___closed__1; +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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_Meta_mkAppM(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___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__1; @@ -254,6 +259,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange___close LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withNewMCtxDepthImp___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14; extern lean_object* l_Lean_Meta_instInhabitedDefaultInstances; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__21; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -263,7 +269,6 @@ static lean_object* l_Lean_Elab_Term_elabForIn___closed__11; lean_object* l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn(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_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel(lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___closed__1; @@ -290,13 +295,14 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn___closed__4; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy(lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Op_elabBinRelCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange___closed__7; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____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_Elab_Term_elabForIn___closed__23; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange(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*); @@ -306,10 +312,9 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___la static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__10; static lean_object* l_Lean_Elab_Term_elabForIn___lambda__1___closed__6; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact(lean_object*); -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__2; LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact(lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__2___closed__1; static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__4; @@ -318,10 +323,9 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_isUnkn static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27_declRange___closed__2; static lean_object* l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__1___closed__1; static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore___lambda__3___closed__2; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange___closed__5; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__11; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange___closed__4; static lean_object* l_Lean_Elab_Term_elabForIn_x27___closed__5; lean_object* l_Lean_Meta_mkDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -338,17 +342,16 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange___closed__3; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Term_Op_elabBinRelCore___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_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_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_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__2; static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__1; double l_Float_ofScientific(lean_object*, uint8_t, lean_object*); lean_object* l_Lean_Elab_Term_ensureHasType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabForIn___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_Syntax_node4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact___closed__2; static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___closed__3; @@ -360,7 +363,6 @@ static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn_x27___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe___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_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_BinOpKind_noConfusion___rarg___lambda__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5; @@ -379,6 +381,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at_Lean_Elab_Term_Op_ela lean_object* l_Lean_Meta_isTypeApp_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(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_Extra_0__Lean_Elab_Term_Op_BinOpKind_toCtorIdx(uint8_t); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_getResetInfoTrees___at_Lean_Elab_Term_withDeclName___spec__3___rarg(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; static lean_object* l_Lean_throwUnknownConstant___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processUnOp___spec__1___closed__3; @@ -391,8 +394,7 @@ lean_object* l_Lean_Meta_coerceSimple_x3f(lean_object*, lean_object*, lean_objec LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange___closed__5; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13; +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__7; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkBinOp___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -409,6 +411,7 @@ static lean_object* l_Lean_Elab_Term_elabForIn___closed__20; LEAN_EXPORT uint8_t l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_hasUncomparable___default; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_max_x3f___default; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processLeaf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15; 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___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange___closed__6; static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn___closed__5; @@ -418,6 +421,7 @@ uint8_t l_Lean_Syntax_isNone(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_declRange___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_throwForInFailure___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___boxed(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_Elab_Term_elabForIn___closed__1; @@ -432,6 +436,7 @@ lean_object* l_Lean_Elab_Term_withInfoContext_x27(lean_object*, lean_object*, le LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange(lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange___closed__2; lean_object* l_Lean_Elab_Term_mkConst(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863_(lean_object*); static lean_object* l_Lean_Elab_Term_elabForIn___closed__15; static lean_object* l_Lean_Elab_Term_elabForIn___closed__2; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__1; @@ -440,7 +445,6 @@ lean_object* l_Lean_Elab_Term_mkCoe(lean_object*, lean_object*, lean_object*, le static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange___closed__2; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy___closed__1; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; static lean_object* l_List_forIn_loop___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__2___lambda__1___closed__1; @@ -458,10 +462,10 @@ static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___cl static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange___closed__3; static lean_object* l_Lean_Elab_Term_Op_elabDefaultOrNonempty___closed__1; static lean_object* l_Lean_Elab_Term_elabForIn___closed__18; -static lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___closed__1; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_x27(lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3(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_elabForIn___closed__6; lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -469,9 +473,9 @@ LEAN_EXPORT uint8_t l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_beqBinOpKin static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp___closed__5; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRange___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze(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_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__19; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExprCore___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___lambda__3___boxed(lean_object**); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___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*); @@ -481,13 +485,13 @@ static lean_object* l_Lean_Elab_Term_Op_elabBinRelCore_toBoolIfNecessary___close 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*); static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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_elabForIn___closed__16; static lean_object* l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRange___closed__2; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___closed__2; -static lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1; +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toExpr___closed__20; -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Term_elabForIn_declRange___closed__3; static lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_applyCoe_go___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_processBinOp___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -5714,7 +5718,7 @@ lean_inc(x_10); x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); -x_12 = 1; +x_12 = 0; x_13 = 0; x_14 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_12, x_13, x_2, x_3, x_4, x_5, x_6, x_7, x_11); if (lean_obj_tag(x_14) == 0) @@ -6197,7 +6201,70 @@ return x_37; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___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, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_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) { +_start: +{ +lean_object* x_11; +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) +{ +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; +} +else +{ +uint8_t 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); +lean_dec(x_2); +x_15 = !lean_is_exclusive(x_11); +if (x_15 == 0) +{ +return x_11; +} +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_EXPORT lean_object* l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg), 10, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___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, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; lean_object* x_12; @@ -6249,15 +6316,15 @@ return x_20; } } } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg___boxed), 10, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg___boxed), 10, 0); return x_2; } } -static lean_object* _init_l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1() { +static lean_object* _init_l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1() { _start: { lean_object* x_1; @@ -6265,11 +6332,11 @@ x_1 = l_Lean_inheritedTraceOptions; return x_1; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; uint8_t x_12; -x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1; +x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1; x_11 = lean_st_ref_get(x_10, x_9); x_12 = !lean_is_exclusive(x_11); if (x_12 == 0) @@ -6302,7 +6369,7 @@ return x_22; } } } -static double _init_l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1() { +static double _init_l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1() { _start: { lean_object* x_1; uint8_t x_2; double x_3; @@ -6312,7 +6379,7 @@ x_3 = l_Float_ofScientific(x_1, x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; 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; @@ -6334,7 +6401,7 @@ if (x_18 == 0) { lean_object* x_19; double x_20; uint8_t x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; x_19 = lean_ctor_get(x_16, 3); -x_20 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1; +x_20 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1; x_21 = 0; x_22 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__5; x_23 = lean_alloc_ctor(0, 2, 17); @@ -6396,7 +6463,7 @@ lean_inc(x_37); lean_inc(x_36); lean_inc(x_35); lean_dec(x_16); -x_42 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1; +x_42 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1; x_43 = 0; x_44 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__5; x_45 = lean_alloc_ctor(0, 2, 17); @@ -6446,15 +6513,141 @@ return x_55; } } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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_10; -x_10 = l_Lean_Meta_isExprDefEqGuarded(x_1, x_2, x_5, x_6, x_7, x_8, x_9); -return x_10; +uint8_t x_11; +x_11 = !lean_is_exclusive(x_6); +if (x_11 == 0) +{ +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_6, 0); +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) +{ +uint8_t x_14; lean_object* x_15; +x_14 = 1; +lean_ctor_set_uint8(x_12, 4, x_14); +x_15 = l_Lean_Meta_isExprDefEqGuarded(x_1, x_2, x_6, x_7, x_8, x_9, x_10); +return x_15; +} +else +{ +uint8_t x_16; uint8_t x_17; uint8_t x_18; uint8_t x_19; uint8_t x_20; uint8_t x_21; uint8_t x_22; uint8_t x_23; uint8_t x_24; uint8_t x_25; uint8_t x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; lean_object* x_30; +x_16 = lean_ctor_get_uint8(x_12, 0); +x_17 = lean_ctor_get_uint8(x_12, 1); +x_18 = lean_ctor_get_uint8(x_12, 2); +x_19 = lean_ctor_get_uint8(x_12, 3); +x_20 = lean_ctor_get_uint8(x_12, 5); +x_21 = lean_ctor_get_uint8(x_12, 6); +x_22 = lean_ctor_get_uint8(x_12, 7); +x_23 = lean_ctor_get_uint8(x_12, 8); +x_24 = lean_ctor_get_uint8(x_12, 9); +x_25 = lean_ctor_get_uint8(x_12, 10); +x_26 = lean_ctor_get_uint8(x_12, 11); +x_27 = lean_ctor_get_uint8(x_12, 12); +lean_dec(x_12); +x_28 = 1; +x_29 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_29, 0, x_16); +lean_ctor_set_uint8(x_29, 1, x_17); +lean_ctor_set_uint8(x_29, 2, x_18); +lean_ctor_set_uint8(x_29, 3, x_19); +lean_ctor_set_uint8(x_29, 4, x_28); +lean_ctor_set_uint8(x_29, 5, x_20); +lean_ctor_set_uint8(x_29, 6, x_21); +lean_ctor_set_uint8(x_29, 7, x_22); +lean_ctor_set_uint8(x_29, 8, x_23); +lean_ctor_set_uint8(x_29, 9, x_24); +lean_ctor_set_uint8(x_29, 10, x_25); +lean_ctor_set_uint8(x_29, 11, x_26); +lean_ctor_set_uint8(x_29, 12, x_27); +lean_ctor_set(x_6, 0, x_29); +x_30 = l_Lean_Meta_isExprDefEqGuarded(x_1, x_2, x_6, x_7, x_8, x_9, x_10); +return x_30; +} } +else +{ +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; uint8_t x_39; uint8_t x_40; uint8_t x_41; uint8_t x_42; uint8_t x_43; uint8_t x_44; uint8_t x_45; uint8_t x_46; uint8_t x_47; uint8_t x_48; uint8_t x_49; uint8_t x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_31 = lean_ctor_get(x_6, 0); +x_32 = lean_ctor_get(x_6, 1); +x_33 = lean_ctor_get(x_6, 2); +x_34 = lean_ctor_get(x_6, 3); +x_35 = lean_ctor_get(x_6, 4); +x_36 = lean_ctor_get(x_6, 5); +x_37 = lean_ctor_get_uint8(x_6, sizeof(void*)*6); +x_38 = lean_ctor_get_uint8(x_6, sizeof(void*)*6 + 1); +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_6); +x_39 = lean_ctor_get_uint8(x_31, 0); +x_40 = lean_ctor_get_uint8(x_31, 1); +x_41 = lean_ctor_get_uint8(x_31, 2); +x_42 = lean_ctor_get_uint8(x_31, 3); +x_43 = lean_ctor_get_uint8(x_31, 5); +x_44 = lean_ctor_get_uint8(x_31, 6); +x_45 = lean_ctor_get_uint8(x_31, 7); +x_46 = lean_ctor_get_uint8(x_31, 8); +x_47 = lean_ctor_get_uint8(x_31, 9); +x_48 = lean_ctor_get_uint8(x_31, 10); +x_49 = lean_ctor_get_uint8(x_31, 11); +x_50 = lean_ctor_get_uint8(x_31, 12); +if (lean_is_exclusive(x_31)) { + x_51 = x_31; +} else { + lean_dec_ref(x_31); + x_51 = lean_box(0); } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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) { +x_52 = 1; +if (lean_is_scalar(x_51)) { + x_53 = lean_alloc_ctor(0, 0, 13); +} else { + x_53 = x_51; +} +lean_ctor_set_uint8(x_53, 0, x_39); +lean_ctor_set_uint8(x_53, 1, x_40); +lean_ctor_set_uint8(x_53, 2, x_41); +lean_ctor_set_uint8(x_53, 3, x_42); +lean_ctor_set_uint8(x_53, 4, x_52); +lean_ctor_set_uint8(x_53, 5, x_43); +lean_ctor_set_uint8(x_53, 6, x_44); +lean_ctor_set_uint8(x_53, 7, x_45); +lean_ctor_set_uint8(x_53, 8, x_46); +lean_ctor_set_uint8(x_53, 9, x_47); +lean_ctor_set_uint8(x_53, 10, x_48); +lean_ctor_set_uint8(x_53, 11, x_49); +lean_ctor_set_uint8(x_53, 12, x_50); +x_54 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_54, 0, x_53); +lean_ctor_set(x_54, 1, x_32); +lean_ctor_set(x_54, 2, x_33); +lean_ctor_set(x_54, 3, x_34); +lean_ctor_set(x_54, 4, x_35); +lean_ctor_set(x_54, 5, x_36); +lean_ctor_set_uint8(x_54, sizeof(void*)*6, x_37); +lean_ctor_set_uint8(x_54, sizeof(void*)*6 + 1, x_38); +x_55 = l_Lean_Meta_isExprDefEqGuarded(x_1, x_2, x_54, x_7, x_8, x_9, x_10); +return x_55; +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; lean_object* x_11; +x_10 = lean_box(x_1); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -6530,6 +6723,14 @@ return x_29; static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1() { _start: { +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2___boxed), 9, 0); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2() { +_start: +{ lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___regBuiltin_Lean_Elab_Term_elabForIn___closed__1; x_2 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_toTree_go___closed__1; @@ -6537,15 +6738,15 @@ x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2() { +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2___boxed), 9, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3___boxed), 9, 0); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3() { +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4() { _start: { lean_object* x_1; @@ -6553,16 +6754,16 @@ x_1 = lean_mk_string_from_bytes("uncomparable types: ", 20); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4() { +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3; +x_1 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5() { +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6() { _start: { lean_object* x_1; @@ -6570,11 +6771,11 @@ x_1 = lean_mk_string_from_bytes(", ", 2); return x_1; } } -static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6() { +static lean_object* _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5; +x_1 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } @@ -6723,16 +6924,18 @@ lean_dec(x_23); x_48 = !lean_is_exclusive(x_25); if (x_48 == 0) { -lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; lean_object* x_53; +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; x_49 = lean_ctor_get(x_25, 0); lean_inc(x_20); lean_inc(x_49); -x_50 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 9, 2); +x_50 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 10, 2); lean_closure_set(x_50, 0, x_49); lean_closure_set(x_50, 1, x_20); -x_51 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_51, 0, x_50); -x_52 = 0; +x_51 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_52 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg), 10, 2); +lean_closure_set(x_52, 0, x_50); +lean_closure_set(x_52, 1, x_51); +x_53 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -6740,20 +6943,20 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_53 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(x_51, x_52, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); -if (lean_obj_tag(x_53) == 0) +x_54 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg(x_52, x_53, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +if (lean_obj_tag(x_54) == 0) { -lean_object* x_54; uint8_t x_55; -x_54 = lean_ctor_get(x_53, 0); -lean_inc(x_54); -x_55 = lean_unbox(x_54); -lean_dec(x_54); -if (x_55 == 0) +lean_object* x_55; uint8_t x_56; +x_55 = lean_ctor_get(x_54, 0); +lean_inc(x_55); +x_56 = lean_unbox(x_55); +lean_dec(x_55); +if (x_56 == 0) { -lean_object* x_56; lean_object* x_57; -x_56 = lean_ctor_get(x_53, 1); -lean_inc(x_56); -lean_dec(x_53); +lean_object* x_57; lean_object* x_58; +x_57 = lean_ctor_get(x_54, 1); +lean_inc(x_57); +lean_dec(x_54); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -6762,20 +6965,20 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_49); lean_inc(x_20); -x_57 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_20, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_56); -if (lean_obj_tag(x_57) == 0) +x_58 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_20, x_49, x_3, x_4, x_5, x_6, x_7, x_8, x_57); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_58; uint8_t x_59; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_unbox(x_58); -lean_dec(x_58); -if (x_59 == 0) +lean_object* x_59; uint8_t x_60; +x_59 = lean_ctor_get(x_58, 0); +lean_inc(x_59); +x_60 = lean_unbox(x_59); +lean_dec(x_59); +if (x_60 == 0) { -lean_object* x_60; lean_object* x_61; -x_60 = lean_ctor_get(x_57, 1); -lean_inc(x_60); -lean_dec(x_57); +lean_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -6784,73 +6987,73 @@ lean_inc(x_4); lean_inc(x_3); lean_inc(x_20); lean_inc(x_49); -x_61 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_49, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_60); -if (lean_obj_tag(x_61) == 0) +x_62 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_49, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_61); +if (lean_obj_tag(x_62) == 0) { -lean_object* x_62; uint8_t x_63; -x_62 = lean_ctor_get(x_61, 0); -lean_inc(x_62); -x_63 = lean_unbox(x_62); -lean_dec(x_62); -if (x_63 == 0) +lean_object* x_63; uint8_t x_64; +x_63 = lean_ctor_get(x_62, 0); +lean_inc(x_63); +x_64 = lean_unbox(x_63); +lean_dec(x_63); +if (x_64 == 0) { -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_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; lean_free_object(x_25); -x_64 = lean_ctor_get(x_61, 1); -lean_inc(x_64); -lean_dec(x_61); -x_65 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; -x_66 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(x_65, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_64); -x_67 = lean_ctor_get(x_66, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_66, 1); +x_65 = lean_ctor_get(x_62, 1); +lean_inc(x_65); +lean_dec(x_62); +x_66 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; +x_67 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_66, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_65); +x_68 = lean_ctor_get(x_67, 0); lean_inc(x_68); -lean_dec(x_66); -x_69 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; -x_70 = lean_unbox(x_67); +x_69 = lean_ctor_get(x_67, 1); +lean_inc(x_69); lean_dec(x_67); -if (x_70 == 0) +x_70 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3; +x_71 = lean_unbox(x_68); +lean_dec(x_68); +if (x_71 == 0) { -lean_object* x_71; lean_object* x_72; +lean_object* x_72; lean_object* x_73; lean_dec(x_49); lean_dec(x_20); -x_71 = lean_box(0); -x_72 = lean_apply_9(x_69, x_71, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_68); -return x_72; -} -else -{ -lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_73 = l_Lean_MessageData_ofExpr(x_49); -x_74 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4; -x_75 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_75, 0, x_74); -lean_ctor_set(x_75, 1, x_73); -x_76 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6; -x_77 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_77, 0, x_75); -lean_ctor_set(x_77, 1, x_76); -x_78 = l_Lean_MessageData_ofExpr(x_20); -x_79 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_79, 0, x_77); -lean_ctor_set(x_79, 1, x_78); -x_80 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; -x_81 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_81, 0, x_79); -lean_ctor_set(x_81, 1, x_80); -x_82 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_65, x_81, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_68); -x_83 = lean_ctor_get(x_82, 0); -lean_inc(x_83); -x_84 = lean_ctor_get(x_82, 1); +x_72 = lean_box(0); +x_73 = lean_apply_9(x_70, x_72, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_69); +return x_73; +} +else +{ +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_74 = l_Lean_MessageData_ofExpr(x_49); +x_75 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5; +x_76 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_76, 0, x_75); +lean_ctor_set(x_76, 1, x_74); +x_77 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7; +x_78 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_78, 0, x_76); +lean_ctor_set(x_78, 1, x_77); +x_79 = l_Lean_MessageData_ofExpr(x_20); +x_80 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; +x_82 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_82, 0, x_80); +lean_ctor_set(x_82, 1, x_81); +x_83 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(x_66, x_82, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_69); +x_84 = lean_ctor_get(x_83, 0); lean_inc(x_84); -lean_dec(x_82); -x_85 = lean_apply_9(x_69, x_83, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_84); -return x_85; +x_85 = lean_ctor_get(x_83, 1); +lean_inc(x_85); +lean_dec(x_83); +x_86 = lean_apply_9(x_70, x_84, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_85); +return x_86; } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; uint8_t x_90; +lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; uint8_t x_91; lean_dec(x_49); lean_dec(x_8); lean_dec(x_7); @@ -6858,83 +7061,83 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_86 = lean_ctor_get(x_61, 1); -lean_inc(x_86); -lean_dec(x_61); -x_87 = lean_st_ref_take(x_2, x_86); -x_88 = lean_ctor_get(x_87, 0); -lean_inc(x_88); -x_89 = lean_ctor_get(x_87, 1); +x_87 = lean_ctor_get(x_62, 1); +lean_inc(x_87); +lean_dec(x_62); +x_88 = lean_st_ref_take(x_2, x_87); +x_89 = lean_ctor_get(x_88, 0); lean_inc(x_89); -lean_dec(x_87); +x_90 = lean_ctor_get(x_88, 1); +lean_inc(x_90); +lean_dec(x_88); lean_ctor_set(x_25, 0, x_20); -x_90 = !lean_is_exclusive(x_88); -if (x_90 == 0) +x_91 = !lean_is_exclusive(x_89); +if (x_91 == 0) { -lean_object* x_91; lean_object* x_92; uint8_t x_93; -x_91 = lean_ctor_get(x_88, 0); -lean_dec(x_91); -lean_ctor_set(x_88, 0, x_25); -x_92 = lean_st_ref_set(x_2, x_88, x_89); +lean_object* x_92; lean_object* x_93; uint8_t x_94; +x_92 = lean_ctor_get(x_89, 0); +lean_dec(x_92); +lean_ctor_set(x_89, 0, x_25); +x_93 = lean_st_ref_set(x_2, x_89, x_90); lean_dec(x_2); -x_93 = !lean_is_exclusive(x_92); -if (x_93 == 0) +x_94 = !lean_is_exclusive(x_93); +if (x_94 == 0) { -lean_object* x_94; lean_object* x_95; -x_94 = lean_ctor_get(x_92, 0); -lean_dec(x_94); -x_95 = lean_box(0); -lean_ctor_set(x_92, 0, x_95); -return x_92; +lean_object* x_95; lean_object* x_96; +x_95 = lean_ctor_get(x_93, 0); +lean_dec(x_95); +x_96 = lean_box(0); +lean_ctor_set(x_93, 0, x_96); +return x_93; } else { -lean_object* x_96; lean_object* x_97; lean_object* x_98; -x_96 = lean_ctor_get(x_92, 1); -lean_inc(x_96); -lean_dec(x_92); -x_97 = lean_box(0); -x_98 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_98, 0, x_97); -lean_ctor_set(x_98, 1, x_96); -return x_98; +lean_object* x_97; lean_object* x_98; lean_object* x_99; +x_97 = lean_ctor_get(x_93, 1); +lean_inc(x_97); +lean_dec(x_93); +x_98 = lean_box(0); +x_99 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_99, 0, x_98); +lean_ctor_set(x_99, 1, x_97); +return x_99; } } else { -uint8_t x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_99 = lean_ctor_get_uint8(x_88, sizeof(void*)*1); -lean_dec(x_88); -x_100 = lean_alloc_ctor(0, 1, 1); -lean_ctor_set(x_100, 0, x_25); -lean_ctor_set_uint8(x_100, sizeof(void*)*1, x_99); -x_101 = lean_st_ref_set(x_2, x_100, x_89); +uint8_t x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; +x_100 = lean_ctor_get_uint8(x_89, sizeof(void*)*1); +lean_dec(x_89); +x_101 = lean_alloc_ctor(0, 1, 1); +lean_ctor_set(x_101, 0, x_25); +lean_ctor_set_uint8(x_101, sizeof(void*)*1, x_100); +x_102 = lean_st_ref_set(x_2, x_101, x_90); lean_dec(x_2); -x_102 = lean_ctor_get(x_101, 1); -lean_inc(x_102); -if (lean_is_exclusive(x_101)) { - lean_ctor_release(x_101, 0); - lean_ctor_release(x_101, 1); - x_103 = x_101; +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_104 = x_102; } else { - lean_dec_ref(x_101); - x_103 = lean_box(0); + lean_dec_ref(x_102); + x_104 = lean_box(0); } -x_104 = lean_box(0); -if (lean_is_scalar(x_103)) { - x_105 = lean_alloc_ctor(0, 2, 0); +x_105 = lean_box(0); +if (lean_is_scalar(x_104)) { + x_106 = lean_alloc_ctor(0, 2, 0); } else { - x_105 = x_103; + x_106 = x_104; } -lean_ctor_set(x_105, 0, x_104); -lean_ctor_set(x_105, 1, x_102); -return x_105; +lean_ctor_set(x_106, 0, x_105); +lean_ctor_set(x_106, 1, x_103); +return x_106; } } } else { -uint8_t x_106; +uint8_t x_107; lean_free_object(x_25); lean_dec(x_49); lean_dec(x_20); @@ -6945,29 +7148,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_106 = !lean_is_exclusive(x_61); -if (x_106 == 0) +x_107 = !lean_is_exclusive(x_62); +if (x_107 == 0) { -return x_61; +return x_62; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_61, 0); -x_108 = lean_ctor_get(x_61, 1); +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_62, 0); +x_109 = lean_ctor_get(x_62, 1); +lean_inc(x_109); lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_61); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +lean_dec(x_62); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; } } } else { -uint8_t x_110; +uint8_t x_111; lean_free_object(x_25); lean_dec(x_49); lean_dec(x_20); @@ -6978,33 +7181,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_110 = !lean_is_exclusive(x_57); -if (x_110 == 0) +x_111 = !lean_is_exclusive(x_58); +if (x_111 == 0) { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_57, 0); -lean_dec(x_111); -x_112 = lean_box(0); -lean_ctor_set(x_57, 0, x_112); -return x_57; +lean_object* x_112; lean_object* x_113; +x_112 = lean_ctor_get(x_58, 0); +lean_dec(x_112); +x_113 = lean_box(0); +lean_ctor_set(x_58, 0, x_113); +return x_58; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_57, 1); -lean_inc(x_113); -lean_dec(x_57); -x_114 = lean_box(0); -x_115 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_115, 0, x_114); -lean_ctor_set(x_115, 1, x_113); -return x_115; +lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_114 = lean_ctor_get(x_58, 1); +lean_inc(x_114); +lean_dec(x_58); +x_115 = lean_box(0); +x_116 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_116, 0, x_115); +lean_ctor_set(x_116, 1, x_114); +return x_116; } } } else { -uint8_t x_116; +uint8_t x_117; lean_free_object(x_25); lean_dec(x_49); lean_dec(x_20); @@ -7015,29 +7218,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_116 = !lean_is_exclusive(x_57); -if (x_116 == 0) +x_117 = !lean_is_exclusive(x_58); +if (x_117 == 0) { -return x_57; +return x_58; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -x_117 = lean_ctor_get(x_57, 0); -x_118 = lean_ctor_get(x_57, 1); +lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_118 = lean_ctor_get(x_58, 0); +x_119 = lean_ctor_get(x_58, 1); +lean_inc(x_119); lean_inc(x_118); -lean_inc(x_117); -lean_dec(x_57); -x_119 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_119, 0, x_117); -lean_ctor_set(x_119, 1, x_118); -return x_119; +lean_dec(x_58); +x_120 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_120, 0, x_118); +lean_ctor_set(x_120, 1, x_119); +return x_120; } } } else { -uint8_t x_120; +uint8_t x_121; lean_free_object(x_25); lean_dec(x_49); lean_dec(x_20); @@ -7048,33 +7251,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_120 = !lean_is_exclusive(x_53); -if (x_120 == 0) -{ -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_53, 0); -lean_dec(x_121); -x_122 = lean_box(0); -lean_ctor_set(x_53, 0, x_122); -return x_53; +x_121 = !lean_is_exclusive(x_54); +if (x_121 == 0) +{ +lean_object* x_122; lean_object* x_123; +x_122 = lean_ctor_get(x_54, 0); +lean_dec(x_122); +x_123 = lean_box(0); +lean_ctor_set(x_54, 0, x_123); +return x_54; } else { -lean_object* x_123; lean_object* x_124; lean_object* x_125; -x_123 = lean_ctor_get(x_53, 1); -lean_inc(x_123); -lean_dec(x_53); -x_124 = lean_box(0); -x_125 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_125, 0, x_124); -lean_ctor_set(x_125, 1, x_123); -return x_125; +lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_124 = lean_ctor_get(x_54, 1); +lean_inc(x_124); +lean_dec(x_54); +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_124); +return x_126; } } } else { -uint8_t x_126; +uint8_t x_127; lean_free_object(x_25); lean_dec(x_49); lean_dec(x_20); @@ -7085,40 +7288,42 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_126 = !lean_is_exclusive(x_53); -if (x_126 == 0) +x_127 = !lean_is_exclusive(x_54); +if (x_127 == 0) { -return x_53; +return x_54; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_53, 0); -x_128 = lean_ctor_get(x_53, 1); +lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_128 = lean_ctor_get(x_54, 0); +x_129 = lean_ctor_get(x_54, 1); +lean_inc(x_129); lean_inc(x_128); -lean_inc(x_127); -lean_dec(x_53); -x_129 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_129, 0, x_127); -lean_ctor_set(x_129, 1, x_128); -return x_129; +lean_dec(x_54); +x_130 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_130, 0, x_128); +lean_ctor_set(x_130, 1, x_129); +return x_130; } } } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; uint8_t x_133; lean_object* x_134; -x_130 = lean_ctor_get(x_25, 0); -lean_inc(x_130); +lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; uint8_t x_135; lean_object* x_136; +x_131 = lean_ctor_get(x_25, 0); +lean_inc(x_131); lean_dec(x_25); lean_inc(x_20); -lean_inc(x_130); -x_131 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 9, 2); -lean_closure_set(x_131, 0, x_130); -lean_closure_set(x_131, 1, x_20); -x_132 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); +lean_inc(x_131); +x_132 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 10, 2); lean_closure_set(x_132, 0, x_131); -x_133 = 0; +lean_closure_set(x_132, 1, x_20); +x_133 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_134 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg), 10, 2); +lean_closure_set(x_134, 0, x_132); +lean_closure_set(x_134, 1, x_133); +x_135 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -7126,42 +7331,42 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_134 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(x_132, x_133, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); -if (lean_obj_tag(x_134) == 0) -{ -lean_object* x_135; uint8_t x_136; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +x_136 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg(x_134, x_135, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_47); +if (lean_obj_tag(x_136) == 0) { -lean_object* x_137; lean_object* x_138; -x_137 = lean_ctor_get(x_134, 1); +lean_object* x_137; uint8_t x_138; +x_137 = lean_ctor_get(x_136, 0); lean_inc(x_137); -lean_dec(x_134); +x_138 = lean_unbox(x_137); +lean_dec(x_137); +if (x_138 == 0) +{ +lean_object* x_139; lean_object* x_140; +x_139 = lean_ctor_get(x_136, 1); +lean_inc(x_139); +lean_dec(x_136); 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_130); +lean_inc(x_131); lean_inc(x_20); -x_138 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_20, x_130, x_3, x_4, x_5, x_6, x_7, x_8, x_137); -if (lean_obj_tag(x_138) == 0) -{ -lean_object* x_139; uint8_t x_140; -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_unbox(x_139); -lean_dec(x_139); -if (x_140 == 0) +x_140 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_20, x_131, x_3, x_4, x_5, x_6, x_7, x_8, x_139); +if (lean_obj_tag(x_140) == 0) { -lean_object* x_141; lean_object* x_142; -x_141 = lean_ctor_get(x_138, 1); +lean_object* x_141; uint8_t x_142; +x_141 = lean_ctor_get(x_140, 0); lean_inc(x_141); -lean_dec(x_138); +x_142 = lean_unbox(x_141); +lean_dec(x_141); +if (x_142 == 0) +{ +lean_object* x_143; lean_object* x_144; +x_143 = lean_ctor_get(x_140, 1); +lean_inc(x_143); +lean_dec(x_140); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -7169,133 +7374,133 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_20); -lean_inc(x_130); -x_142 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_130, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_141); -if (lean_obj_tag(x_142) == 0) -{ -lean_object* x_143; uint8_t x_144; -x_143 = lean_ctor_get(x_142, 0); -lean_inc(x_143); -x_144 = lean_unbox(x_143); -lean_dec(x_143); -if (x_144 == 0) +lean_inc(x_131); +x_144 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_131, x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_143); +if (lean_obj_tag(x_144) == 0) { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; uint8_t x_151; -x_145 = lean_ctor_get(x_142, 1); +lean_object* x_145; uint8_t x_146; +x_145 = lean_ctor_get(x_144, 0); lean_inc(x_145); -lean_dec(x_142); -x_146 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; -x_147 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(x_146, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_145); -x_148 = lean_ctor_get(x_147, 0); -lean_inc(x_148); -x_149 = lean_ctor_get(x_147, 1); -lean_inc(x_149); -lean_dec(x_147); -x_150 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; -x_151 = lean_unbox(x_148); -lean_dec(x_148); -if (x_151 == 0) -{ -lean_object* x_152; lean_object* x_153; -lean_dec(x_130); +x_146 = lean_unbox(x_145); +lean_dec(x_145); +if (x_146 == 0) +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; uint8_t x_153; +x_147 = lean_ctor_get(x_144, 1); +lean_inc(x_147); +lean_dec(x_144); +x_148 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; +x_149 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_148, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_147); +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 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3; +x_153 = lean_unbox(x_150); +lean_dec(x_150); +if (x_153 == 0) +{ +lean_object* x_154; lean_object* x_155; +lean_dec(x_131); lean_dec(x_20); -x_152 = lean_box(0); -x_153 = lean_apply_9(x_150, x_152, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_149); -return x_153; +x_154 = lean_box(0); +x_155 = lean_apply_9(x_152, x_154, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_151); +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; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; -x_154 = l_Lean_MessageData_ofExpr(x_130); -x_155 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4; -x_156 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_156, 0, x_155); -lean_ctor_set(x_156, 1, x_154); -x_157 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6; +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_156 = l_Lean_MessageData_ofExpr(x_131); +x_157 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5; x_158 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_158, 0, x_156); -lean_ctor_set(x_158, 1, x_157); -x_159 = l_Lean_MessageData_ofExpr(x_20); +lean_ctor_set(x_158, 0, x_157); +lean_ctor_set(x_158, 1, x_156); +x_159 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7; x_160 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_160, 0, x_158); lean_ctor_set(x_160, 1, x_159); -x_161 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; +x_161 = l_Lean_MessageData_ofExpr(x_20); x_162 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_162, 0, x_160); lean_ctor_set(x_162, 1, x_161); -x_163 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_146, x_162, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_149); -x_164 = lean_ctor_get(x_163, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_163, 1); -lean_inc(x_165); -lean_dec(x_163); -x_166 = lean_apply_9(x_150, x_164, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_165); -return x_166; +x_163 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; +x_164 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +x_165 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(x_148, x_164, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_151); +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_apply_9(x_152, x_166, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_167); +return x_168; } } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; -lean_dec(x_130); +lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; uint8_t x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; +lean_dec(x_131); 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_167 = lean_ctor_get(x_142, 1); -lean_inc(x_167); -lean_dec(x_142); -x_168 = lean_st_ref_take(x_2, x_167); -x_169 = lean_ctor_get(x_168, 0); +x_169 = lean_ctor_get(x_144, 1); lean_inc(x_169); -x_170 = lean_ctor_get(x_168, 1); -lean_inc(x_170); -lean_dec(x_168); -x_171 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_171, 0, x_20); -x_172 = lean_ctor_get_uint8(x_169, sizeof(void*)*1); -if (lean_is_exclusive(x_169)) { - lean_ctor_release(x_169, 0); - x_173 = x_169; +lean_dec(x_144); +x_170 = lean_st_ref_take(x_2, x_169); +x_171 = lean_ctor_get(x_170, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_170, 1); +lean_inc(x_172); +lean_dec(x_170); +x_173 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_173, 0, x_20); +x_174 = lean_ctor_get_uint8(x_171, sizeof(void*)*1); +if (lean_is_exclusive(x_171)) { + lean_ctor_release(x_171, 0); + x_175 = x_171; } else { - lean_dec_ref(x_169); - x_173 = lean_box(0); + lean_dec_ref(x_171); + x_175 = lean_box(0); } -if (lean_is_scalar(x_173)) { - x_174 = lean_alloc_ctor(0, 1, 1); +if (lean_is_scalar(x_175)) { + x_176 = lean_alloc_ctor(0, 1, 1); } else { - x_174 = x_173; + x_176 = x_175; } -lean_ctor_set(x_174, 0, x_171); -lean_ctor_set_uint8(x_174, sizeof(void*)*1, x_172); -x_175 = lean_st_ref_set(x_2, x_174, x_170); +lean_ctor_set(x_176, 0, x_173); +lean_ctor_set_uint8(x_176, sizeof(void*)*1, x_174); +x_177 = lean_st_ref_set(x_2, x_176, x_172); lean_dec(x_2); -x_176 = lean_ctor_get(x_175, 1); -lean_inc(x_176); -if (lean_is_exclusive(x_175)) { - lean_ctor_release(x_175, 0); - lean_ctor_release(x_175, 1); - x_177 = x_175; +x_178 = lean_ctor_get(x_177, 1); +lean_inc(x_178); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_179 = x_177; } else { - lean_dec_ref(x_175); - x_177 = lean_box(0); + lean_dec_ref(x_177); + x_179 = lean_box(0); } -x_178 = lean_box(0); -if (lean_is_scalar(x_177)) { - x_179 = lean_alloc_ctor(0, 2, 0); +x_180 = lean_box(0); +if (lean_is_scalar(x_179)) { + x_181 = lean_alloc_ctor(0, 2, 0); } else { - x_179 = x_177; + x_181 = x_179; } -lean_ctor_set(x_179, 0, x_178); -lean_ctor_set(x_179, 1, x_176); -return x_179; +lean_ctor_set(x_181, 0, x_180); +lean_ctor_set(x_181, 1, x_178); +return x_181; } } else { -lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; -lean_dec(x_130); +lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; +lean_dec(x_131); lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7304,32 +7509,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_180 = lean_ctor_get(x_142, 0); -lean_inc(x_180); -x_181 = lean_ctor_get(x_142, 1); -lean_inc(x_181); -if (lean_is_exclusive(x_142)) { - lean_ctor_release(x_142, 0); - lean_ctor_release(x_142, 1); - x_182 = x_142; +x_182 = lean_ctor_get(x_144, 0); +lean_inc(x_182); +x_183 = lean_ctor_get(x_144, 1); +lean_inc(x_183); +if (lean_is_exclusive(x_144)) { + lean_ctor_release(x_144, 0); + lean_ctor_release(x_144, 1); + x_184 = x_144; } else { - lean_dec_ref(x_142); - x_182 = lean_box(0); + lean_dec_ref(x_144); + x_184 = lean_box(0); } -if (lean_is_scalar(x_182)) { - x_183 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_184)) { + x_185 = lean_alloc_ctor(1, 2, 0); } else { - x_183 = x_182; + x_185 = x_184; } -lean_ctor_set(x_183, 0, x_180); -lean_ctor_set(x_183, 1, x_181); -return x_183; +lean_ctor_set(x_185, 0, x_182); +lean_ctor_set(x_185, 1, x_183); +return x_185; } } else { -lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; -lean_dec(x_130); +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; +lean_dec(x_131); lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7338,31 +7543,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_184 = lean_ctor_get(x_138, 1); -lean_inc(x_184); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_185 = x_138; +x_186 = lean_ctor_get(x_140, 1); +lean_inc(x_186); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_187 = x_140; } else { - lean_dec_ref(x_138); - x_185 = lean_box(0); + lean_dec_ref(x_140); + x_187 = lean_box(0); } -x_186 = lean_box(0); -if (lean_is_scalar(x_185)) { - x_187 = lean_alloc_ctor(0, 2, 0); +x_188 = lean_box(0); +if (lean_is_scalar(x_187)) { + x_189 = lean_alloc_ctor(0, 2, 0); } else { - x_187 = x_185; + x_189 = x_187; } -lean_ctor_set(x_187, 0, x_186); -lean_ctor_set(x_187, 1, x_184); -return x_187; +lean_ctor_set(x_189, 0, x_188); +lean_ctor_set(x_189, 1, x_186); +return x_189; } } else { -lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; -lean_dec(x_130); +lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; +lean_dec(x_131); lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7371,32 +7576,32 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_188 = lean_ctor_get(x_138, 0); -lean_inc(x_188); -x_189 = lean_ctor_get(x_138, 1); -lean_inc(x_189); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_190 = x_138; +x_190 = lean_ctor_get(x_140, 0); +lean_inc(x_190); +x_191 = lean_ctor_get(x_140, 1); +lean_inc(x_191); +if (lean_is_exclusive(x_140)) { + lean_ctor_release(x_140, 0); + lean_ctor_release(x_140, 1); + x_192 = x_140; } else { - lean_dec_ref(x_138); - x_190 = lean_box(0); + lean_dec_ref(x_140); + x_192 = lean_box(0); } -if (lean_is_scalar(x_190)) { - x_191 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_192)) { + x_193 = lean_alloc_ctor(1, 2, 0); } else { - x_191 = x_190; + x_193 = x_192; } -lean_ctor_set(x_191, 0, x_188); -lean_ctor_set(x_191, 1, x_189); -return x_191; +lean_ctor_set(x_193, 0, x_190); +lean_ctor_set(x_193, 1, x_191); +return x_193; } } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; -lean_dec(x_130); +lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; +lean_dec(x_131); lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7405,31 +7610,31 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_192 = lean_ctor_get(x_134, 1); -lean_inc(x_192); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_193 = x_134; +x_194 = lean_ctor_get(x_136, 1); +lean_inc(x_194); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_195 = x_136; } else { - lean_dec_ref(x_134); - x_193 = lean_box(0); + lean_dec_ref(x_136); + x_195 = lean_box(0); } -x_194 = lean_box(0); -if (lean_is_scalar(x_193)) { - x_195 = lean_alloc_ctor(0, 2, 0); +x_196 = lean_box(0); +if (lean_is_scalar(x_195)) { + x_197 = lean_alloc_ctor(0, 2, 0); } else { - x_195 = x_193; + x_197 = x_195; } -lean_ctor_set(x_195, 0, x_194); -lean_ctor_set(x_195, 1, x_192); -return x_195; +lean_ctor_set(x_197, 0, x_196); +lean_ctor_set(x_197, 1, x_194); +return x_197; } } else { -lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_130); +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +lean_dec(x_131); lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7438,33 +7643,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_196 = lean_ctor_get(x_134, 0); -lean_inc(x_196); -x_197 = lean_ctor_get(x_134, 1); -lean_inc(x_197); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_198 = x_134; +x_198 = lean_ctor_get(x_136, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_136, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_136)) { + lean_ctor_release(x_136, 0); + lean_ctor_release(x_136, 1); + x_200 = x_136; } else { - lean_dec_ref(x_134); - x_198 = lean_box(0); + lean_dec_ref(x_136); + x_200 = lean_box(0); } -if (lean_is_scalar(x_198)) { - x_199 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_200)) { + x_201 = lean_alloc_ctor(1, 2, 0); } else { - x_199 = x_198; + x_201 = x_200; } -lean_ctor_set(x_199, 0, x_196); -lean_ctor_set(x_199, 1, x_197); -return x_199; +lean_ctor_set(x_201, 0, x_198); +lean_ctor_set(x_201, 1, x_199); +return x_201; } } } } else { -lean_object* x_200; +lean_object* x_202; lean_dec(x_20); lean_dec(x_8); lean_dec(x_7); @@ -7473,109 +7678,111 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_200 = lean_box(0); -lean_ctor_set(x_18, 0, x_200); +x_202 = lean_box(0); +lean_ctor_set(x_18, 0, x_202); return x_18; } } else { -lean_object* x_201; lean_object* x_202; uint8_t x_203; -x_201 = lean_ctor_get(x_18, 0); -x_202 = lean_ctor_get(x_18, 1); -lean_inc(x_202); -lean_inc(x_201); +lean_object* x_203; lean_object* x_204; uint8_t x_205; +x_203 = lean_ctor_get(x_18, 0); +x_204 = lean_ctor_get(x_18, 1); +lean_inc(x_204); +lean_inc(x_203); lean_dec(x_18); -x_203 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_isUnknown(x_201); -if (x_203 == 0) +x_205 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_isUnknown(x_203); +if (x_205 == 0) { -lean_object* x_204; lean_object* x_205; lean_object* x_206; -x_204 = lean_st_ref_get(x_2, x_202); -x_205 = lean_ctor_get(x_204, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_205, 0); -lean_inc(x_206); -lean_dec(x_205); -if (lean_obj_tag(x_206) == 0) +lean_object* x_206; lean_object* x_207; lean_object* x_208; +x_206 = lean_st_ref_get(x_2, x_204); +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +x_208 = lean_ctor_get(x_207, 0); +lean_inc(x_208); +lean_dec(x_207); +if (lean_obj_tag(x_208) == 0) { -lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; lean_object* x_211; uint8_t x_212; lean_object* x_213; lean_object* x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; +lean_object* x_209; lean_object* x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; uint8_t x_214; lean_object* x_215; lean_object* x_216; lean_object* x_217; lean_object* x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_207 = lean_ctor_get(x_204, 1); -lean_inc(x_207); -lean_dec(x_204); -x_208 = lean_st_ref_take(x_2, x_207); -x_209 = lean_ctor_get(x_208, 0); +x_209 = lean_ctor_get(x_206, 1); lean_inc(x_209); -x_210 = lean_ctor_get(x_208, 1); -lean_inc(x_210); -lean_dec(x_208); -x_211 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_211, 0, x_201); -x_212 = lean_ctor_get_uint8(x_209, sizeof(void*)*1); -if (lean_is_exclusive(x_209)) { - lean_ctor_release(x_209, 0); - x_213 = x_209; +lean_dec(x_206); +x_210 = lean_st_ref_take(x_2, x_209); +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_ctor_get(x_210, 1); +lean_inc(x_212); +lean_dec(x_210); +x_213 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_213, 0, x_203); +x_214 = lean_ctor_get_uint8(x_211, sizeof(void*)*1); +if (lean_is_exclusive(x_211)) { + lean_ctor_release(x_211, 0); + x_215 = x_211; } else { - lean_dec_ref(x_209); - x_213 = lean_box(0); + lean_dec_ref(x_211); + x_215 = lean_box(0); } -if (lean_is_scalar(x_213)) { - x_214 = lean_alloc_ctor(0, 1, 1); +if (lean_is_scalar(x_215)) { + x_216 = lean_alloc_ctor(0, 1, 1); } else { - x_214 = x_213; + x_216 = x_215; } -lean_ctor_set(x_214, 0, x_211); -lean_ctor_set_uint8(x_214, sizeof(void*)*1, x_212); -x_215 = lean_st_ref_set(x_2, x_214, x_210); +lean_ctor_set(x_216, 0, x_213); +lean_ctor_set_uint8(x_216, sizeof(void*)*1, x_214); +x_217 = lean_st_ref_set(x_2, x_216, x_212); lean_dec(x_2); -x_216 = lean_ctor_get(x_215, 1); -lean_inc(x_216); -if (lean_is_exclusive(x_215)) { - lean_ctor_release(x_215, 0); - lean_ctor_release(x_215, 1); - x_217 = x_215; +x_218 = lean_ctor_get(x_217, 1); +lean_inc(x_218); +if (lean_is_exclusive(x_217)) { + lean_ctor_release(x_217, 0); + lean_ctor_release(x_217, 1); + x_219 = x_217; } else { - lean_dec_ref(x_215); - x_217 = lean_box(0); + lean_dec_ref(x_217); + x_219 = lean_box(0); } -x_218 = lean_box(0); -if (lean_is_scalar(x_217)) { - x_219 = lean_alloc_ctor(0, 2, 0); +x_220 = lean_box(0); +if (lean_is_scalar(x_219)) { + x_221 = lean_alloc_ctor(0, 2, 0); } else { - x_219 = x_217; + x_221 = x_219; } -lean_ctor_set(x_219, 0, x_218); -lean_ctor_set(x_219, 1, x_216); -return x_219; +lean_ctor_set(x_221, 0, x_220); +lean_ctor_set(x_221, 1, x_218); +return x_221; } else { -lean_object* x_220; lean_object* x_221; lean_object* x_222; lean_object* x_223; lean_object* x_224; uint8_t x_225; lean_object* x_226; -x_220 = lean_ctor_get(x_204, 1); -lean_inc(x_220); -lean_dec(x_204); -x_221 = lean_ctor_get(x_206, 0); -lean_inc(x_221); -if (lean_is_exclusive(x_206)) { - lean_ctor_release(x_206, 0); - x_222 = x_206; +lean_object* x_222; lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_object* x_227; uint8_t x_228; lean_object* x_229; +x_222 = lean_ctor_get(x_206, 1); +lean_inc(x_222); +lean_dec(x_206); +x_223 = lean_ctor_get(x_208, 0); +lean_inc(x_223); +if (lean_is_exclusive(x_208)) { + lean_ctor_release(x_208, 0); + x_224 = x_208; } else { - lean_dec_ref(x_206); - x_222 = lean_box(0); + lean_dec_ref(x_208); + x_224 = lean_box(0); } -lean_inc(x_201); -lean_inc(x_221); -x_223 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 9, 2); -lean_closure_set(x_223, 0, x_221); -lean_closure_set(x_223, 1, x_201); -x_224 = lean_alloc_closure((void*)(l_StateRefT_x27_lift___rarg___boxed), 2, 1); -lean_closure_set(x_224, 0, x_223); -x_225 = 0; +lean_inc(x_203); +lean_inc(x_223); +x_225 = lean_alloc_closure((void*)(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed), 10, 2); +lean_closure_set(x_225, 0, x_223); +lean_closure_set(x_225, 1, x_203); +x_226 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_227 = lean_alloc_closure((void*)(l_ReaderT_bind___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg), 10, 2); +lean_closure_set(x_227, 0, x_225); +lean_closure_set(x_227, 1, x_226); +x_228 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -7583,183 +7790,183 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_226 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(x_224, x_225, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_220); -if (lean_obj_tag(x_226) == 0) +x_229 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg(x_227, x_228, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_222); +if (lean_obj_tag(x_229) == 0) { -lean_object* x_227; uint8_t x_228; -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_unbox(x_227); -lean_dec(x_227); -if (x_228 == 0) +lean_object* x_230; uint8_t x_231; +x_230 = lean_ctor_get(x_229, 0); +lean_inc(x_230); +x_231 = lean_unbox(x_230); +lean_dec(x_230); +if (x_231 == 0) { -lean_object* x_229; lean_object* x_230; -x_229 = lean_ctor_get(x_226, 1); -lean_inc(x_229); -lean_dec(x_226); +lean_object* x_232; lean_object* x_233; +x_232 = lean_ctor_get(x_229, 1); +lean_inc(x_232); +lean_dec(x_229); 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_221); -lean_inc(x_201); -x_230 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_201, x_221, x_3, x_4, x_5, x_6, x_7, x_8, x_229); -if (lean_obj_tag(x_230) == 0) +lean_inc(x_223); +lean_inc(x_203); +x_233 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_203, x_223, x_3, x_4, x_5, x_6, x_7, x_8, x_232); +if (lean_obj_tag(x_233) == 0) { -lean_object* x_231; uint8_t x_232; -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_unbox(x_231); -lean_dec(x_231); -if (x_232 == 0) +lean_object* x_234; uint8_t x_235; +x_234 = lean_ctor_get(x_233, 0); +lean_inc(x_234); +x_235 = lean_unbox(x_234); +lean_dec(x_234); +if (x_235 == 0) { -lean_object* x_233; lean_object* x_234; -x_233 = lean_ctor_get(x_230, 1); -lean_inc(x_233); -lean_dec(x_230); +lean_object* x_236; lean_object* x_237; +x_236 = lean_ctor_get(x_233, 1); +lean_inc(x_236); +lean_dec(x_233); 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_201); -lean_inc(x_221); -x_234 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_221, x_201, x_3, x_4, x_5, x_6, x_7, x_8, x_233); -if (lean_obj_tag(x_234) == 0) +lean_inc(x_203); +lean_inc(x_223); +x_237 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe(x_223, x_203, x_3, x_4, x_5, x_6, x_7, x_8, x_236); +if (lean_obj_tag(x_237) == 0) { -lean_object* x_235; uint8_t x_236; -x_235 = lean_ctor_get(x_234, 0); -lean_inc(x_235); -x_236 = lean_unbox(x_235); -lean_dec(x_235); -if (x_236 == 0) +lean_object* x_238; uint8_t x_239; +x_238 = lean_ctor_get(x_237, 0); +lean_inc(x_238); +x_239 = lean_unbox(x_238); +lean_dec(x_238); +if (x_239 == 0) { -lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; uint8_t x_243; -lean_dec(x_222); -x_237 = lean_ctor_get(x_234, 1); -lean_inc(x_237); -lean_dec(x_234); -x_238 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; -x_239 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(x_238, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_237); -x_240 = lean_ctor_get(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; uint8_t x_246; +lean_dec(x_224); +x_240 = lean_ctor_get(x_237, 1); lean_inc(x_240); -x_241 = lean_ctor_get(x_239, 1); -lean_inc(x_241); -lean_dec(x_239); -x_242 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; -x_243 = lean_unbox(x_240); -lean_dec(x_240); -if (x_243 == 0) -{ -lean_object* x_244; lean_object* x_245; -lean_dec(x_221); -lean_dec(x_201); -x_244 = lean_box(0); -x_245 = lean_apply_9(x_242, x_244, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_241); -return x_245; -} -else -{ -lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; -x_246 = l_Lean_MessageData_ofExpr(x_221); -x_247 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__4; -x_248 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_248, 0, x_247); -lean_ctor_set(x_248, 1, x_246); -x_249 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6; -x_250 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_250, 0, x_248); -lean_ctor_set(x_250, 1, x_249); -x_251 = l_Lean_MessageData_ofExpr(x_201); -x_252 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_252, 0, x_250); -lean_ctor_set(x_252, 1, x_251); -x_253 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; -x_254 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_254, 0, x_252); -lean_ctor_set(x_254, 1, x_253); -x_255 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_238, x_254, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_241); -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_apply_9(x_242, x_256, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_257); -return x_258; +lean_dec(x_237); +x_241 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; +x_242 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_241, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_240); +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); +lean_inc(x_244); +lean_dec(x_242); +x_245 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__3; +x_246 = lean_unbox(x_243); +lean_dec(x_243); +if (x_246 == 0) +{ +lean_object* x_247; lean_object* x_248; +lean_dec(x_223); +lean_dec(x_203); +x_247 = lean_box(0); +x_248 = lean_apply_9(x_245, x_247, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_244); +return x_248; +} +else +{ +lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; lean_object* x_260; lean_object* x_261; +x_249 = l_Lean_MessageData_ofExpr(x_223); +x_250 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5; +x_251 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_251, 0, x_250); +lean_ctor_set(x_251, 1, x_249); +x_252 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7; +x_253 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_253, 0, x_251); +lean_ctor_set(x_253, 1, x_252); +x_254 = l_Lean_MessageData_ofExpr(x_203); +x_255 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_255, 0, x_253); +lean_ctor_set(x_255, 1, x_254); +x_256 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_getMonadForIn___closed__6; +x_257 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_257, 0, x_255); +lean_ctor_set(x_257, 1, x_256); +x_258 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(x_241, x_257, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_244); +x_259 = lean_ctor_get(x_258, 0); +lean_inc(x_259); +x_260 = lean_ctor_get(x_258, 1); +lean_inc(x_260); +lean_dec(x_258); +x_261 = lean_apply_9(x_245, x_259, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_260); +return x_261; } } else { -lean_object* x_259; lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; uint8_t x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; lean_object* x_271; -lean_dec(x_221); +lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; uint8_t 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_dec(x_223); 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_259 = lean_ctor_get(x_234, 1); -lean_inc(x_259); -lean_dec(x_234); -x_260 = lean_st_ref_take(x_2, x_259); -x_261 = lean_ctor_get(x_260, 0); -lean_inc(x_261); -x_262 = lean_ctor_get(x_260, 1); +x_262 = lean_ctor_get(x_237, 1); lean_inc(x_262); -lean_dec(x_260); -if (lean_is_scalar(x_222)) { - x_263 = lean_alloc_ctor(1, 1, 0); +lean_dec(x_237); +x_263 = lean_st_ref_take(x_2, x_262); +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); +if (lean_is_scalar(x_224)) { + x_266 = lean_alloc_ctor(1, 1, 0); } else { - x_263 = x_222; + x_266 = x_224; } -lean_ctor_set(x_263, 0, x_201); -x_264 = lean_ctor_get_uint8(x_261, sizeof(void*)*1); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - x_265 = x_261; +lean_ctor_set(x_266, 0, x_203); +x_267 = lean_ctor_get_uint8(x_264, sizeof(void*)*1); +if (lean_is_exclusive(x_264)) { + lean_ctor_release(x_264, 0); + x_268 = x_264; } else { - lean_dec_ref(x_261); - x_265 = lean_box(0); + lean_dec_ref(x_264); + x_268 = lean_box(0); } -if (lean_is_scalar(x_265)) { - x_266 = lean_alloc_ctor(0, 1, 1); +if (lean_is_scalar(x_268)) { + x_269 = lean_alloc_ctor(0, 1, 1); } else { - x_266 = x_265; + x_269 = x_268; } -lean_ctor_set(x_266, 0, x_263); -lean_ctor_set_uint8(x_266, sizeof(void*)*1, x_264); -x_267 = lean_st_ref_set(x_2, x_266, x_262); +lean_ctor_set(x_269, 0, x_266); +lean_ctor_set_uint8(x_269, sizeof(void*)*1, x_267); +x_270 = lean_st_ref_set(x_2, x_269, x_265); lean_dec(x_2); -x_268 = lean_ctor_get(x_267, 1); -lean_inc(x_268); -if (lean_is_exclusive(x_267)) { - lean_ctor_release(x_267, 0); - lean_ctor_release(x_267, 1); - x_269 = x_267; +x_271 = lean_ctor_get(x_270, 1); +lean_inc(x_271); +if (lean_is_exclusive(x_270)) { + lean_ctor_release(x_270, 0); + lean_ctor_release(x_270, 1); + x_272 = x_270; } else { - lean_dec_ref(x_267); - x_269 = lean_box(0); + lean_dec_ref(x_270); + x_272 = lean_box(0); } -x_270 = lean_box(0); -if (lean_is_scalar(x_269)) { - x_271 = lean_alloc_ctor(0, 2, 0); +x_273 = lean_box(0); +if (lean_is_scalar(x_272)) { + x_274 = lean_alloc_ctor(0, 2, 0); } else { - x_271 = x_269; + x_274 = x_272; } -lean_ctor_set(x_271, 0, x_270); -lean_ctor_set(x_271, 1, x_268); -return x_271; +lean_ctor_set(x_274, 0, x_273); +lean_ctor_set(x_274, 1, x_271); +return x_274; } } else { -lean_object* x_272; lean_object* x_273; lean_object* x_274; lean_object* x_275; -lean_dec(x_222); -lean_dec(x_221); -lean_dec(x_201); +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7767,34 +7974,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_272 = lean_ctor_get(x_234, 0); -lean_inc(x_272); -x_273 = lean_ctor_get(x_234, 1); -lean_inc(x_273); -if (lean_is_exclusive(x_234)) { - lean_ctor_release(x_234, 0); - lean_ctor_release(x_234, 1); - x_274 = x_234; +x_275 = lean_ctor_get(x_237, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_237, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_277 = x_237; } else { - lean_dec_ref(x_234); - x_274 = lean_box(0); + lean_dec_ref(x_237); + x_277 = lean_box(0); } -if (lean_is_scalar(x_274)) { - x_275 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_275 = x_274; + x_278 = x_277; } -lean_ctor_set(x_275, 0, x_272); -lean_ctor_set(x_275, 1, x_273); -return x_275; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } else { -lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_object* x_279; -lean_dec(x_222); -lean_dec(x_221); -lean_dec(x_201); +lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7802,33 +8009,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_276 = lean_ctor_get(x_230, 1); -lean_inc(x_276); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_277 = x_230; +x_279 = lean_ctor_get(x_233, 1); +lean_inc(x_279); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_280 = x_233; } else { - lean_dec_ref(x_230); - x_277 = lean_box(0); + lean_dec_ref(x_233); + x_280 = lean_box(0); } -x_278 = lean_box(0); -if (lean_is_scalar(x_277)) { - x_279 = lean_alloc_ctor(0, 2, 0); +x_281 = lean_box(0); +if (lean_is_scalar(x_280)) { + x_282 = lean_alloc_ctor(0, 2, 0); } else { - x_279 = x_277; + x_282 = x_280; } -lean_ctor_set(x_279, 0, x_278); -lean_ctor_set(x_279, 1, x_276); -return x_279; +lean_ctor_set(x_282, 0, x_281); +lean_ctor_set(x_282, 1, x_279); +return x_282; } } else { -lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -lean_dec(x_222); -lean_dec(x_221); -lean_dec(x_201); +lean_object* x_283; lean_object* x_284; lean_object* x_285; lean_object* x_286; +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7836,34 +8043,34 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_280 = lean_ctor_get(x_230, 0); -lean_inc(x_280); -x_281 = lean_ctor_get(x_230, 1); -lean_inc(x_281); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_282 = x_230; +x_283 = lean_ctor_get(x_233, 0); +lean_inc(x_283); +x_284 = lean_ctor_get(x_233, 1); +lean_inc(x_284); +if (lean_is_exclusive(x_233)) { + lean_ctor_release(x_233, 0); + lean_ctor_release(x_233, 1); + x_285 = x_233; } else { - lean_dec_ref(x_230); - x_282 = lean_box(0); + lean_dec_ref(x_233); + x_285 = lean_box(0); } -if (lean_is_scalar(x_282)) { - x_283 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_285)) { + x_286 = lean_alloc_ctor(1, 2, 0); } else { - x_283 = x_282; + x_286 = x_285; } -lean_ctor_set(x_283, 0, x_280); -lean_ctor_set(x_283, 1, x_281); -return x_283; +lean_ctor_set(x_286, 0, x_283); +lean_ctor_set(x_286, 1, x_284); +return x_286; } } else { -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; -lean_dec(x_222); -lean_dec(x_221); -lean_dec(x_201); +lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7871,33 +8078,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_284 = lean_ctor_get(x_226, 1); -lean_inc(x_284); -if (lean_is_exclusive(x_226)) { - lean_ctor_release(x_226, 0); - lean_ctor_release(x_226, 1); - x_285 = x_226; +x_287 = lean_ctor_get(x_229, 1); +lean_inc(x_287); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + lean_ctor_release(x_229, 1); + x_288 = x_229; } else { - lean_dec_ref(x_226); - x_285 = lean_box(0); + lean_dec_ref(x_229); + x_288 = lean_box(0); } -x_286 = lean_box(0); -if (lean_is_scalar(x_285)) { - x_287 = lean_alloc_ctor(0, 2, 0); +x_289 = lean_box(0); +if (lean_is_scalar(x_288)) { + x_290 = lean_alloc_ctor(0, 2, 0); } else { - x_287 = x_285; + x_290 = x_288; } -lean_ctor_set(x_287, 0, x_286); -lean_ctor_set(x_287, 1, x_284); -return x_287; +lean_ctor_set(x_290, 0, x_289); +lean_ctor_set(x_290, 1, x_287); +return x_290; } } else { -lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; -lean_dec(x_222); -lean_dec(x_221); -lean_dec(x_201); +lean_object* x_291; lean_object* x_292; lean_object* x_293; lean_object* x_294; +lean_dec(x_224); +lean_dec(x_223); +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7905,33 +8112,33 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_288 = lean_ctor_get(x_226, 0); -lean_inc(x_288); -x_289 = lean_ctor_get(x_226, 1); -lean_inc(x_289); -if (lean_is_exclusive(x_226)) { - lean_ctor_release(x_226, 0); - lean_ctor_release(x_226, 1); - x_290 = x_226; +x_291 = lean_ctor_get(x_229, 0); +lean_inc(x_291); +x_292 = lean_ctor_get(x_229, 1); +lean_inc(x_292); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + lean_ctor_release(x_229, 1); + x_293 = x_229; } else { - lean_dec_ref(x_226); - x_290 = lean_box(0); + lean_dec_ref(x_229); + x_293 = lean_box(0); } -if (lean_is_scalar(x_290)) { - x_291 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_293)) { + x_294 = lean_alloc_ctor(1, 2, 0); } else { - x_291 = x_290; + x_294 = x_293; } -lean_ctor_set(x_291, 0, x_288); -lean_ctor_set(x_291, 1, x_289); -return x_291; +lean_ctor_set(x_294, 0, x_291); +lean_ctor_set(x_294, 1, x_292); +return x_294; } } } else { -lean_object* x_292; lean_object* x_293; -lean_dec(x_201); +lean_object* x_295; lean_object* x_296; +lean_dec(x_203); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7939,17 +8146,17 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_292 = lean_box(0); -x_293 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_293, 0, x_292); -lean_ctor_set(x_293, 1, x_202); -return x_293; +x_295 = lean_box(0); +x_296 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_296, 0, x_295); +lean_ctor_set(x_296, 1, x_204); +return x_296; } } } else { -uint8_t x_294; +uint8_t x_297; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -7957,69 +8164,69 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_294 = !lean_is_exclusive(x_15); -if (x_294 == 0) +x_297 = !lean_is_exclusive(x_15); +if (x_297 == 0) { return x_15; } else { -lean_object* x_295; lean_object* x_296; lean_object* x_297; -x_295 = lean_ctor_get(x_15, 0); -x_296 = lean_ctor_get(x_15, 1); -lean_inc(x_296); -lean_inc(x_295); +lean_object* x_298; lean_object* x_299; lean_object* x_300; +x_298 = lean_ctor_get(x_15, 0); +x_299 = lean_ctor_get(x_15, 1); +lean_inc(x_299); +lean_inc(x_298); lean_dec(x_15); -x_297 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_297, 0, x_295); -lean_ctor_set(x_297, 1, x_296); -return x_297; +x_300 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_300, 0, x_298); +lean_ctor_set(x_300, 1, x_299); +return x_300; } } } case 1: { -uint8_t x_298; lean_object* x_299; -x_298 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); -x_299 = lean_box(x_298); -switch (lean_obj_tag(x_299)) { +uint8_t x_301; lean_object* x_302; +x_301 = lean_ctor_get_uint8(x_1, sizeof(void*)*4); +x_302 = lean_box(x_301); +switch (lean_obj_tag(x_302)) { case 2: { -lean_object* x_300; lean_object* x_301; -x_300 = lean_ctor_get(x_10, 1); -lean_inc(x_300); -lean_dec(x_10); -x_301 = lean_ctor_get(x_1, 3); -lean_inc(x_301); -lean_dec(x_1); -x_1 = x_301; -x_9 = x_300; -goto _start; -} -case 3: -{ lean_object* x_303; lean_object* x_304; x_303 = lean_ctor_get(x_10, 1); lean_inc(x_303); lean_dec(x_10); -x_304 = lean_ctor_get(x_1, 2); +x_304 = lean_ctor_get(x_1, 3); lean_inc(x_304); lean_dec(x_1); x_1 = x_304; x_9 = x_303; goto _start; } -default: +case 3: { -lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; -lean_dec(x_299); +lean_object* x_306; lean_object* x_307; x_306 = lean_ctor_get(x_10, 1); lean_inc(x_306); lean_dec(x_10); x_307 = lean_ctor_get(x_1, 2); lean_inc(x_307); -x_308 = lean_ctor_get(x_1, 3); -lean_inc(x_308); +lean_dec(x_1); +x_1 = x_307; +x_9 = x_306; +goto _start; +} +default: +{ +lean_object* x_309; lean_object* x_310; lean_object* x_311; lean_object* x_312; +lean_dec(x_302); +x_309 = lean_ctor_get(x_10, 1); +lean_inc(x_309); +lean_dec(x_10); +x_310 = lean_ctor_get(x_1, 2); +lean_inc(x_310); +x_311 = lean_ctor_get(x_1, 3); +lean_inc(x_311); lean_dec(x_1); lean_inc(x_8); lean_inc(x_7); @@ -8028,21 +8235,21 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_309 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go(x_307, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_306); -if (lean_obj_tag(x_309) == 0) +x_312 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go(x_310, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_309); +if (lean_obj_tag(x_312) == 0) { -lean_object* x_310; -x_310 = lean_ctor_get(x_309, 1); -lean_inc(x_310); -lean_dec(x_309); -x_1 = x_308; -x_9 = x_310; +lean_object* x_313; +x_313 = lean_ctor_get(x_312, 1); +lean_inc(x_313); +lean_dec(x_312); +x_1 = x_311; +x_9 = x_313; goto _start; } else { -uint8_t x_312; -lean_dec(x_308); +uint8_t x_315; +lean_dec(x_311); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8050,23 +8257,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_312 = !lean_is_exclusive(x_309); -if (x_312 == 0) +x_315 = !lean_is_exclusive(x_312); +if (x_315 == 0) { -return x_309; +return x_312; } else { -lean_object* x_313; lean_object* x_314; lean_object* x_315; -x_313 = lean_ctor_get(x_309, 0); -x_314 = lean_ctor_get(x_309, 1); -lean_inc(x_314); -lean_inc(x_313); -lean_dec(x_309); -x_315 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_315, 0, x_313); -lean_ctor_set(x_315, 1, x_314); -return x_315; +lean_object* x_316; lean_object* x_317; lean_object* x_318; +x_316 = lean_ctor_get(x_312, 0); +x_317 = lean_ctor_get(x_312, 1); +lean_inc(x_317); +lean_inc(x_316); +lean_dec(x_312); +x_318 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_318, 0, x_316); +lean_ctor_set(x_318, 1, x_317); +return x_318; } } } @@ -8074,35 +8281,35 @@ return x_315; } case 2: { -lean_object* x_316; lean_object* x_317; -x_316 = lean_ctor_get(x_10, 1); -lean_inc(x_316); -lean_dec(x_10); -x_317 = lean_ctor_get(x_1, 2); -lean_inc(x_317); -lean_dec(x_1); -x_1 = x_317; -x_9 = x_316; -goto _start; -} -default: -{ lean_object* x_319; lean_object* x_320; x_319 = lean_ctor_get(x_10, 1); lean_inc(x_319); lean_dec(x_10); -x_320 = lean_ctor_get(x_1, 3); +x_320 = lean_ctor_get(x_1, 2); lean_inc(x_320); lean_dec(x_1); x_1 = x_320; x_9 = x_319; goto _start; } +default: +{ +lean_object* x_322; lean_object* x_323; +x_322 = lean_ctor_get(x_10, 1); +lean_inc(x_322); +lean_dec(x_10); +x_323 = lean_ctor_get(x_1, 3); +lean_inc(x_323); +lean_dec(x_1); +x_1 = x_323; +x_9 = x_322; +goto _start; +} } } else { -uint8_t x_322; +uint8_t x_325; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8111,27 +8318,27 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_322 = !lean_is_exclusive(x_10); -if (x_322 == 0) +x_325 = !lean_is_exclusive(x_10); +if (x_325 == 0) { -lean_object* x_323; lean_object* x_324; -x_323 = lean_ctor_get(x_10, 0); -lean_dec(x_323); -x_324 = lean_box(0); -lean_ctor_set(x_10, 0, x_324); +lean_object* x_326; lean_object* x_327; +x_326 = lean_ctor_get(x_10, 0); +lean_dec(x_326); +x_327 = lean_box(0); +lean_ctor_set(x_10, 0, x_327); return x_10; } else { -lean_object* x_325; lean_object* x_326; lean_object* x_327; -x_325 = lean_ctor_get(x_10, 1); -lean_inc(x_325); +lean_object* x_328; lean_object* x_329; lean_object* x_330; +x_328 = lean_ctor_get(x_10, 1); +lean_inc(x_328); lean_dec(x_10); -x_326 = lean_box(0); -x_327 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_327, 0, x_326); -lean_ctor_set(x_327, 1, x_325); -return x_327; +x_329 = lean_box(0); +x_330 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_330, 0, x_329); +lean_ctor_set(x_330, 1, x_328); +return x_330; } } } @@ -8151,21 +8358,21 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_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_EXPORT lean_object* l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___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) { _start: { uint8_t x_11; lean_object* x_12; x_11 = lean_unbox(x_2); lean_dec(x_2); -x_12 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__2___rarg(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_12 = l_Lean_Meta_withNewMCtxDepth___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___rarg(x_1, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_12; } } -LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { lean_object* x_10; -x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -8176,11 +8383,11 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +LEAN_EXPORT lean_object* l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { lean_object* x_11; -x_11 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -8191,21 +8398,39 @@ lean_dec(x_3); return x_11; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_10; -x_10 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_object* x_11; +x_11 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -return x_10; +return x_11; } } LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_1); +lean_dec(x_1); +x_11 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_11; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ lean_object* x_10; -x_10 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -10608,7 +10833,7 @@ lean_inc(x_20); x_21 = lean_ctor_get(x_19, 1); lean_inc(x_21); lean_dec(x_19); -x_22 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_22 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; x_23 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_22, x_6, x_7, x_8, x_9, x_10, x_11, x_21); x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); @@ -12970,7 +13195,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_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_13 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; x_14 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); @@ -13325,7 +13550,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__ _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(465u); x_2 = lean_unsigned_to_nat(27u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13337,7 +13562,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__ _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(465u); x_2 = lean_unsigned_to_nat(61u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13365,7 +13590,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__ _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(465u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13377,7 +13602,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOp_declRange__ _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(465u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13476,7 +13701,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(440u); +x_1 = lean_unsigned_to_nat(466u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13488,7 +13713,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(440u); +x_1 = lean_unsigned_to_nat(466u); x_2 = lean_unsigned_to_nat(70u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13516,7 +13741,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(440u); +x_1 = lean_unsigned_to_nat(466u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13528,7 +13753,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinOpLazy_declRan _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(440u); +x_1 = lean_unsigned_to_nat(466u); x_2 = lean_unsigned_to_nat(49u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13627,7 +13852,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange _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(467u); x_2 = lean_unsigned_to_nat(29u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13639,7 +13864,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange _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(467u); x_2 = lean_unsigned_to_nat(65u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13667,7 +13892,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange _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(467u); x_2 = lean_unsigned_to_nat(33u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13679,7 +13904,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabLeftact_declRange _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(467u); x_2 = lean_unsigned_to_nat(44u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13778,7 +14003,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(442u); +x_1 = lean_unsigned_to_nat(468u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13790,7 +14015,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(442u); +x_1 = lean_unsigned_to_nat(468u); x_2 = lean_unsigned_to_nat(67u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13818,7 +14043,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(442u); +x_1 = lean_unsigned_to_nat(468u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13830,7 +14055,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabRightact_declRang _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(442u); +x_1 = lean_unsigned_to_nat(468u); x_2 = lean_unsigned_to_nat(46u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13929,7 +14154,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(443u); +x_1 = lean_unsigned_to_nat(469u); x_2 = lean_unsigned_to_nat(26u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13941,7 +14166,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(443u); +x_1 = lean_unsigned_to_nat(469u); x_2 = lean_unsigned_to_nat(59u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13969,7 +14194,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(443u); +x_1 = lean_unsigned_to_nat(469u); x_2 = lean_unsigned_to_nat(30u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13981,7 +14206,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabUnOp_declRange___ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(443u); +x_1 = lean_unsigned_to_nat(469u); x_2 = lean_unsigned_to_nat(38u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15543,7 +15768,7 @@ return x_19; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; lean_object* x_29; +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_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_12); x_20 = lean_ctor_get(x_15, 1); lean_inc(x_20); @@ -15563,14 +15788,13 @@ lean_closure_set(x_27, 2, x_2); lean_closure_set(x_27, 3, x_21); lean_closure_set(x_27, 4, x_26); lean_closure_set(x_27, 5, x_3); -x_28 = 1; -x_29 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_27, x_28, x_14, x_4, x_5, x_6, x_7, x_8, x_9, x_20); -return x_29; +x_28 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_20); +return x_28; } } else { -uint8_t x_30; +uint8_t x_29; lean_dec(x_12); lean_dec(x_9); lean_dec(x_8); @@ -15580,23 +15804,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_15); -if (x_30 == 0) +x_29 = !lean_is_exclusive(x_15); +if (x_29 == 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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_15, 0); +x_31 = lean_ctor_get(x_15, 1); lean_inc(x_31); +lean_inc(x_30); lean_dec(x_15); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +x_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; } } } @@ -15739,7 +15963,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(526u); +x_1 = lean_unsigned_to_nat(551u); x_2 = lean_unsigned_to_nat(28u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15751,7 +15975,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(526u); +x_1 = lean_unsigned_to_nat(551u); x_2 = lean_unsigned_to_nat(77u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15779,7 +16003,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(526u); +x_1 = lean_unsigned_to_nat(551u); x_2 = lean_unsigned_to_nat(32u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15791,7 +16015,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRel_declRange_ _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(526u); +x_1 = lean_unsigned_to_nat(551u); x_2 = lean_unsigned_to_nat(42u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15911,7 +16135,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(528u); +x_1 = lean_unsigned_to_nat(553u); x_2 = lean_unsigned_to_nat(36u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15923,7 +16147,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(528u); +x_1 = lean_unsigned_to_nat(553u); x_2 = lean_unsigned_to_nat(90u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15951,7 +16175,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(528u); +x_1 = lean_unsigned_to_nat(553u); x_2 = lean_unsigned_to_nat(40u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -15963,7 +16187,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabBinRelNoProp_decl _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(528u); +x_1 = lean_unsigned_to_nat(553u); x_2 = lean_unsigned_to_nat(56u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16714,7 +16938,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(531u); +x_1 = lean_unsigned_to_nat(556u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16726,7 +16950,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(546u); +x_1 = lean_unsigned_to_nat(571u); x_2 = lean_unsigned_to_nat(34u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16754,7 +16978,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(531u); +x_1 = lean_unsigned_to_nat(556u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16766,7 +16990,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(531u); +x_1 = lean_unsigned_to_nat(556u); x_2 = lean_unsigned_to_nat(25u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -16812,7 +17036,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -16822,37 +17046,37 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1; x_2 = l___regBuiltin_Lean_Elab_Term_elabForIn___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2; x_2 = l_Lean_Elab_Term_elabForIn___closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__3; x_2 = l___regBuiltin_Lean_Elab_Term_Op_elabBinOp___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5() { _start: { lean_object* x_1; @@ -16860,17 +17084,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7() { _start: { lean_object* x_1; @@ -16878,37 +17102,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8; x_2 = l_Lean_Elab_Term_elabForIn___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9; x_2 = l___regBuiltin_Lean_Elab_Term_elabForIn___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11() { _start: { lean_object* x_1; @@ -16916,17 +17140,17 @@ x_1 = lean_mk_string_from_bytes("Extra", 5); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13() { _start: { lean_object* x_1; @@ -16934,33 +17158,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12; -x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13; +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12; +x_2 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14; -x_2 = lean_unsigned_to_nat(6839u); +x_1 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14; +x_2 = lean_unsigned_to_nat(6863u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863_(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_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1; +x_2 = l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15; +x_4 = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -17228,9 +17452,9 @@ lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasCoe___c l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_max_x3f___default = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_max_x3f___default(); lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_max_x3f___default); l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_hasUncomparable___default = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_AnalyzeResult_hasUncomparable___default(); -l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1 = _init_l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1(); -lean_mark_persistent(l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__3___closed__1); -l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1(); +l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1 = _init_l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1(); +lean_mark_persistent(l_Lean_isTracingEnabledFor___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__4___closed__1); +l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1 = _init_l_Lean_addTrace___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___spec__5___closed__1(); l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1 = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__1); l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2 = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__2(); @@ -17243,6 +17467,8 @@ l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5 = _init_ lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__5); l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6 = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6(); lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__6); +l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7 = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7(); +lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_analyze_go___closed__7); l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkUnOp___closed__1 = _init_l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkUnOp___closed__1(); lean_mark_persistent(l___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_mkUnOp___closed__1); l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__1___closed__1 = _init_l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_Extra_0__Lean_Elab_Term_Op_hasHeterogeneousDefaultInstances___spec__1___closed__1(); @@ -17561,37 +17787,37 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_decl if (builtin) {res = l___regBuiltin_Lean_Elab_Term_Op_elabDefaultOrNonempty_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__1); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__2); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__3); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__4); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__5); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__6); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__7); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__8); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__9); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__10); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__11); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__12); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__13); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__14); -l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839____closed__15); -if (builtin) {res = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6839_(lean_io_mk_world()); +}l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__1); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__2); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__3 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__3); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__4); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__5); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__6); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__7); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__8); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__9); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__10); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__11); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__12); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__13); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__14); +l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15 = _init_l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863____closed__15); +if (builtin) {res = l_Lean_Elab_Term_Op_initFn____x40_Lean_Elab_Extra___hyg_6863_(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/Inductive.c b/stage0/stdlib/Lean/Elab/Inductive.c index 66e13c646c9a..6cafbe853374 100644 --- a/stage0/stdlib/Lean/Elab/Inductive.c +++ b/stage0/stdlib/Lean/Elab/Inductive.c @@ -1947,134 +1947,135 @@ return x_22; LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_10; lean_object* x_11; -x_10 = 0; +uint8_t x_10; uint8_t x_11; lean_object* x_12; +x_10 = 1; +x_11 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_11 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_10, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_11) == 0) +x_12 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_12) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_12, 1); +lean_inc(x_13); +lean_dec(x_12); +x_14 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = l_Lean_Elab_Term_addAutoBoundImplicits(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_15 = l_Lean_Elab_Term_addAutoBoundImplicits(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_13); lean_dec(x_4); -if (lean_obj_tag(x_14) == 0) +if (lean_obj_tag(x_15) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); +lean_object* x_16; lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_15, 0); lean_inc(x_16); -lean_dec(x_14); -x_17 = 1; +x_17 = lean_ctor_get(x_15, 1); +lean_inc(x_17); +lean_dec(x_15); x_18 = 1; -lean_inc(x_15); -x_19 = l_Lean_Meta_mkForallFVars(x_15, x_1, x_10, x_17, x_18, x_5, x_6, x_7, x_8, x_16); +x_19 = 1; +lean_inc(x_16); +x_20 = l_Lean_Meta_mkForallFVars(x_16, x_1, x_11, x_18, x_19, x_5, x_6, x_7, x_8, x_17); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_19) == 0) +if (lean_obj_tag(x_20) == 0) { -uint8_t x_20; -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = lean_array_get_size(x_15); -lean_dec(x_15); -x_23 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -lean_ctor_set(x_19, 0, x_23); -return x_19; +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_20, 0); +x_23 = lean_array_get_size(x_16); +lean_dec(x_16); +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_20, 0, x_24); +return x_20; } else { -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_19, 0); -x_25 = lean_ctor_get(x_19, 1); +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_20, 0); +x_26 = lean_ctor_get(x_20, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_19); -x_26 = lean_array_get_size(x_15); -lean_dec(x_15); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_24); -lean_ctor_set(x_27, 1, x_26); +lean_dec(x_20); +x_27 = lean_array_get_size(x_16); +lean_dec(x_16); 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; +lean_ctor_set(x_28, 0, x_25); +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_26); +return x_29; } } else { -uint8_t x_29; -lean_dec(x_15); -x_29 = !lean_is_exclusive(x_19); -if (x_29 == 0) +uint8_t x_30; +lean_dec(x_16); +x_30 = !lean_is_exclusive(x_20); +if (x_30 == 0) { -return x_19; +return x_20; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_19, 0); -x_31 = lean_ctor_get(x_19, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_20, 0); +x_32 = lean_ctor_get(x_20, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); -lean_dec(x_19); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +lean_dec(x_20); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } else { -uint8_t x_33; +uint8_t x_34; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_14); -if (x_33 == 0) +x_34 = !lean_is_exclusive(x_15); +if (x_34 == 0) { -return x_14; +return x_15; } 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_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_15, 0); +x_36 = lean_ctor_get(x_15, 1); +lean_inc(x_36); 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_dec(x_15); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } } else { -uint8_t x_37; +uint8_t x_38; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -2082,23 +2083,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_37 = !lean_is_exclusive(x_11); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_12); +if (x_38 == 0) { -return x_11; +return x_12; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_11, 0); -x_39 = lean_ctor_get(x_11, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_12, 0); +x_40 = lean_ctor_get(x_12, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_11); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_12); +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; } } } @@ -2429,7 +2430,7 @@ x_11 = lean_ctor_get(x_1, 7); lean_inc(x_11); if (lean_obj_tag(x_11) == 0) { -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; +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; x_12 = l_Lean_Meta_mkFreshLevelMVar(x_6, x_7, x_8, x_9, x_10); x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); @@ -2437,29 +2438,30 @@ x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); x_15 = l_Lean_Expr_sort___override(x_13); -x_16 = 0; +x_16 = 1; +x_17 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_16, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_16, x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); -x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__2___boxed), 11, 2); -lean_closure_set(x_19, 0, x_1); -lean_closure_set(x_19, 1, x_2); -x_20 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_15, x_19, x_4, x_5, x_6, x_7, x_8, x_9, x_18); -return x_20; +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); +x_20 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__2___boxed), 11, 2); +lean_closure_set(x_20, 0, x_1); +lean_closure_set(x_20, 1, x_2); +x_21 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_15, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +return x_21; } else { -uint8_t x_21; +uint8_t x_22; lean_dec(x_15); lean_dec(x_9); lean_dec(x_8); @@ -2470,61 +2472,61 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_17); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_17; +return x_18; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_17, 0); -x_23 = lean_ctor_get(x_17, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_18, 0); +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_17); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_dec(x_18); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_11, 0); -lean_inc(x_25); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_11, 0); +lean_inc(x_26); lean_dec(x_11); -x_26 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__4), 8, 1); -lean_closure_set(x_26, 0, x_25); +x_27 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__4), 8, 1); +lean_closure_set(x_27, 0, x_26); 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_27 = l_Lean_Elab_Term_withAutoBoundImplicit___rarg(x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_Elab_Term_withAutoBoundImplicit___rarg(x_27, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_28) == 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_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_ctor_get(x_28, 0); +x_30 = lean_ctor_get(x_28, 1); lean_inc(x_30); lean_dec(x_28); -x_31 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6), 11, 2); -lean_closure_set(x_31, 0, x_1); -lean_closure_set(x_31, 1, x_2); -x_32 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_30, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_29); -return x_32; +x_31 = lean_ctor_get(x_29, 0); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_alloc_closure((void*)(l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6), 11, 2); +lean_closure_set(x_32, 0, x_1); +lean_closure_set(x_32, 1, x_2); +x_33 = l_Lean_Elab_Term_addAutoBoundImplicits_x27___rarg(x_3, x_31, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_30); +return x_33; } else { -uint8_t x_33; +uint8_t x_34; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -2534,23 +2536,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_33 = !lean_is_exclusive(x_27); -if (x_33 == 0) +x_34 = !lean_is_exclusive(x_28); +if (x_34 == 0) { -return x_27; +return x_28; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_27, 0); -x_35 = lean_ctor_get(x_27, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_28, 0); +x_36 = lean_ctor_get(x_28, 1); +lean_inc(x_36); lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_27); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; +lean_dec(x_28); +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; } } } @@ -8104,7 +8106,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_el _start: { uint8_t x_14; uint8_t x_15; lean_object* x_16; -x_14 = 1; +x_14 = 0; x_15 = 0; lean_inc(x_12); lean_inc(x_11); @@ -8834,101 +8836,102 @@ return x_2; LEAN_EXPORT lean_object* l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { -uint8_t x_12; lean_object* x_13; -x_12 = 0; +uint8_t x_12; uint8_t x_13; lean_object* x_14; +x_12 = 1; +x_13 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_13 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_12, x_12, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_13) == 0) +x_14 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_14; lean_object* x_15; -x_14 = lean_ctor_get(x_13, 1); -lean_inc(x_14); -lean_dec(x_13); +lean_object* x_15; lean_object* x_16; +x_15 = lean_ctor_get(x_14, 1); +lean_inc(x_15); +lean_dec(x_14); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_5); -x_15 = l_Lean_Elab_Term_addAutoBoundImplicits(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Lean_Elab_Term_addAutoBoundImplicits(x_1, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -lean_inc(x_16); -x_18 = lean_alloc_closure((void*)(l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__1___boxed), 2, 1); -lean_closure_set(x_18, 0, x_16); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +lean_inc(x_17); +x_19 = lean_alloc_closure((void*)(l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__1___boxed), 2, 1); +lean_closure_set(x_19, 0, x_17); lean_inc(x_4); -x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +x_20 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_19); -x_22 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; -x_23 = l_Lean_Elab_Term_collectUnassignedMVars(x_20, x_22, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_21); -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -x_25 = lean_ctor_get(x_23, 1); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = l_Lean_Elab_Command_instInhabitedCtorView___closed__1; +x_24 = l_Lean_Elab_Term_collectUnassignedMVars(x_21, x_23, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_22); +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__3; -x_27 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_26, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_unbox(x_28); -lean_dec(x_28); -if (x_29 == 0) +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_Command_initFn____x40_Lean_Elab_Inductive___hyg_5____closed__3; +x_28 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(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_unbox(x_29); +lean_dec(x_29); +if (x_30 == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_box(0); -x_32 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__3(x_24, x_16, x_4, x_2, x_26, x_3, x_31, x_5, x_6, x_7, x_8, x_9, x_10, x_30); -return x_32; +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_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__3(x_25, x_17, x_4, x_2, x_27, x_3, x_32, x_5, x_6, x_7, x_8, x_9, x_10, x_31); +return x_33; } 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; -x_33 = lean_ctor_get(x_27, 1); -lean_inc(x_33); -lean_dec(x_27); -lean_inc(x_24); -x_34 = lean_array_to_list(lean_box(0), x_24); -x_35 = lean_box(0); -x_36 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_34, x_35); -x_37 = l_Lean_MessageData_ofList(x_36); -lean_dec(x_36); -x_38 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__4___closed__2; -x_39 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_37); -x_40 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6___closed__5; -x_41 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_41, 0, x_39); -lean_ctor_set(x_41, 1, x_40); -x_42 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_26, x_41, x_5, x_6, x_7, x_8, x_9, x_10, x_33); -x_43 = lean_ctor_get(x_42, 0); -lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +lean_inc(x_25); +x_35 = lean_array_to_list(lean_box(0), x_25); +x_36 = lean_box(0); +x_37 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_35, x_36); +x_38 = l_Lean_MessageData_ofList(x_37); +lean_dec(x_37); +x_39 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__4___closed__2; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6___closed__5; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_27, x_42, x_5, x_6, x_7, x_8, x_9, x_10, x_34); +x_44 = lean_ctor_get(x_43, 0); lean_inc(x_44); -lean_dec(x_42); -x_45 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__3(x_24, x_16, x_4, x_2, x_26, x_3, x_43, x_5, x_6, x_7, x_8, x_9, x_10, x_44); -return x_45; +x_45 = lean_ctor_get(x_43, 1); +lean_inc(x_45); +lean_dec(x_43); +x_46 = l_List_mapM_loop___at___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabCtors___spec__2___lambda__3(x_25, x_17, x_4, x_2, x_27, x_3, x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_45); +return x_46; } } else { -uint8_t x_46; +uint8_t x_47; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8938,29 +8941,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_46 = !lean_is_exclusive(x_15); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_16); +if (x_47 == 0) { -return x_15; +return x_16; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_15, 0); -x_48 = lean_ctor_get(x_15, 1); +lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_16, 0); +x_49 = lean_ctor_get(x_16, 1); +lean_inc(x_49); lean_inc(x_48); -lean_inc(x_47); -lean_dec(x_15); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +lean_dec(x_16); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } } else { -uint8_t x_50; +uint8_t x_51; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -8971,23 +8974,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_50 = !lean_is_exclusive(x_13); -if (x_50 == 0) +x_51 = !lean_is_exclusive(x_14); +if (x_51 == 0) { -return x_13; +return x_14; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_13, 0); -x_52 = lean_ctor_get(x_13, 1); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +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_inc(x_51); -lean_dec(x_13); -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_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); +return x_54; } } } @@ -31264,27 +31267,28 @@ lean_dec(x_3); lean_dec(x_2); if (lean_obj_tag(x_22) == 0) { -lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; +lean_object* x_23; lean_object* x_24; uint8_t x_25; uint8_t x_26; lean_object* x_27; x_23 = lean_ctor_get(x_22, 0); lean_inc(x_23); x_24 = lean_ctor_get(x_22, 1); lean_inc(x_24); lean_dec(x_22); -x_25 = 0; +x_25 = 1; +x_26 = 0; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_26 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_25, x_25, x_12, x_13, x_14, x_15, x_16, x_17, x_24); -if (lean_obj_tag(x_26) == 0) +x_27 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_25, x_26, x_12, x_13, x_14, x_15, x_16, x_17, x_24); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 1); -lean_inc(x_27); -lean_dec(x_26); -x_28 = lean_array_get_size(x_4); +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_array_get_size(x_4); lean_dec(x_4); lean_inc(x_17); lean_inc(x_16); @@ -31292,68 +31296,68 @@ lean_inc(x_15); lean_inc(x_14); lean_inc(x_5); lean_inc(x_23); -x_29 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams(x_28, x_23, x_5, x_14, x_15, x_16, x_17, x_27); -if (lean_obj_tag(x_29) == 0) +x_30 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_fixedIndicesToParams(x_29, x_23, x_5, x_14, x_15, x_16, x_17, x_28); +if (lean_obj_tag(x_30) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_31 = lean_ctor_get(x_30, 0); lean_inc(x_31); -lean_dec(x_29); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); lean_inc(x_6); -x_32 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_31); -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) +x_33 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_6, x_12, x_13, x_14, x_15, x_16, x_17, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); +lean_dec(x_34); +if (x_35 == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_dec(x_6); -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_box(0); -x_37 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__3(x_23, x_30, x_7, x_8, x_1, x_5, x_9, x_10, x_36, x_12, x_13, x_14, x_15, x_16, x_17, x_35); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_box(0); +x_38 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__3(x_23, x_31, x_7, x_8, x_1, x_5, x_9, x_10, x_37, x_12, x_13, x_14, x_15, x_16, x_17, x_36); lean_dec(x_10); -return x_37; +return x_38; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_38 = lean_ctor_get(x_32, 1); -lean_inc(x_38); -lean_dec(x_32); -lean_inc(x_30); -x_39 = l___private_Init_Data_Repr_0__Nat_reprFast(x_30); -x_40 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_40, 0, x_39); -x_41 = lean_alloc_ctor(0, 1, 0); +lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; +x_39 = lean_ctor_get(x_33, 1); +lean_inc(x_39); +lean_dec(x_33); +lean_inc(x_31); +x_40 = l___private_Init_Data_Repr_0__Nat_reprFast(x_31); +x_41 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_41, 0, x_40); -x_42 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__4___closed__2; -x_43 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_43, 0, x_42); -lean_ctor_set(x_43, 1, x_41); -x_44 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6___closed__5; -x_45 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_6, x_45, x_12, x_13, x_14, x_15, x_16, x_17, x_38); -x_47 = lean_ctor_get(x_46, 0); -lean_inc(x_47); -x_48 = lean_ctor_get(x_46, 1); +x_42 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_42, 0, x_41); +x_43 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__4___closed__2; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_elabHeaderAux___lambda__6___closed__5; +x_46 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +x_47 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_6, x_46, x_12, x_13, x_14, x_15, x_16, x_17, x_39); +x_48 = lean_ctor_get(x_47, 0); lean_inc(x_48); -lean_dec(x_46); -x_49 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__3(x_23, x_30, x_7, x_8, x_1, x_5, x_9, x_10, x_47, x_12, x_13, x_14, x_15, x_16, x_17, x_48); +x_49 = lean_ctor_get(x_47, 1); +lean_inc(x_49); lean_dec(x_47); +x_50 = l___private_Lean_Elab_Inductive_0__Lean_Elab_Command_mkInductiveDecl___lambda__3(x_23, x_31, x_7, x_8, x_1, x_5, x_9, x_10, x_48, x_12, x_13, x_14, x_15, x_16, x_17, x_49); +lean_dec(x_48); lean_dec(x_10); -return x_49; +return x_50; } } else { -uint8_t x_50; +uint8_t x_51; lean_dec(x_23); lean_dec(x_17); lean_dec(x_16); @@ -31367,29 +31371,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_1); -x_50 = !lean_is_exclusive(x_29); -if (x_50 == 0) +x_51 = !lean_is_exclusive(x_30); +if (x_51 == 0) { -return x_29; +return x_30; } else { -lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_51 = lean_ctor_get(x_29, 0); -x_52 = lean_ctor_get(x_29, 1); +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_30, 0); +x_53 = lean_ctor_get(x_30, 1); +lean_inc(x_53); lean_inc(x_52); -lean_inc(x_51); -lean_dec(x_29); -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_dec(x_30); +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_54; +uint8_t x_55; lean_dec(x_23); lean_dec(x_17); lean_dec(x_16); @@ -31404,29 +31408,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_54 = !lean_is_exclusive(x_26); -if (x_54 == 0) +x_55 = !lean_is_exclusive(x_27); +if (x_55 == 0) { -return x_26; +return x_27; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_26, 0); -x_56 = lean_ctor_get(x_26, 1); +lean_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_27, 0); +x_57 = lean_ctor_get(x_27, 1); +lean_inc(x_57); lean_inc(x_56); -lean_inc(x_55); -lean_dec(x_26); -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_dec(x_27); +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_58; +uint8_t x_59; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -31440,23 +31444,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_58 = !lean_is_exclusive(x_22); -if (x_58 == 0) +x_59 = !lean_is_exclusive(x_22); +if (x_59 == 0) { return x_22; } 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_object* x_60; lean_object* x_61; lean_object* x_62; +x_60 = lean_ctor_get(x_22, 0); +x_61 = lean_ctor_get(x_22, 1); +lean_inc(x_61); 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; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/InheritDoc.c b/stage0/stdlib/Lean/Elab/InheritDoc.c index 60c9303e646a..99edebd73eb2 100644 --- a/stage0/stdlib/Lean/Elab/InheritDoc.c +++ b/stage0/stdlib/Lean/Elab/InheritDoc.c @@ -56,6 +56,7 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4_(lean static lean_object* l_Lean_addDocString___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__5___lambda__1___closed__2; static lean_object* l_Lean_addDocString___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__5___lambda__1___closed__4; static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____closed__14; +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__3___closed__9; LEAN_EXPORT lean_object* l_Lean_log___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__5___closed__1; @@ -104,7 +105,6 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_initFn____x40_L static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____closed__5; static lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____spec__6(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); @@ -2244,7 +2244,7 @@ LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Elab_InheritDoc___hyg_4____la { uint8_t x_8; uint8_t x_9; x_8 = 0; -x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_8); +x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_8); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; uint8_t x_12; diff --git a/stage0/stdlib/Lean/Elab/Match.c b/stage0/stdlib/Lean/Elab/Match.c index 6efad700c53e..d15c7906a139 100644 --- a/stage0/stdlib/Lean/Elab/Match.c +++ b/stage0/stdlib/Lean/Elab/Match.c @@ -696,7 +696,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_ToDepElimPatter extern lean_object* l_Lean_Elab_abortTermExceptionId; LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltViews_updateMatchType___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_HashMap_insert___at_Lean_instantiateExprMVars___spec__3(lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Match___hyg_18632____closed__14; lean_object* l_Lean_CollectFVars_main(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchAltView___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -16726,18 +16726,18 @@ lean_inc(x_10); x_35 = lean_whnf(x_32, x_10, x_11, x_12, x_13, x_14); if (x_34 == 0) { -lean_object* x_218; lean_object* x_219; -x_218 = l_Lean_instInhabitedSyntax; -x_219 = l___private_Init_GetElem_0__outOfBounds___rarg(x_218); -x_36 = x_219; -goto block_217; +lean_object* x_222; lean_object* x_223; +x_222 = l_Lean_instInhabitedSyntax; +x_223 = l___private_Init_GetElem_0__outOfBounds___rarg(x_222); +x_36 = x_223; +goto block_221; } else { -lean_object* x_220; -x_220 = lean_array_fget(x_1, x_4); -x_36 = x_220; -goto block_217; +lean_object* x_224; +x_224 = lean_array_fget(x_1, x_4); +x_36 = x_224; +goto block_221; } block_31: { @@ -16793,7 +16793,7 @@ x_14 = x_21; goto _start; } } -block_217: +block_221: { if (lean_obj_tag(x_35) == 0) { @@ -16832,14 +16832,14 @@ lean_closure_set(x_55, 3, x_54); lean_closure_set(x_55, 4, x_51); x_56 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); lean_closure_set(x_56, 0, x_55); -x_57 = 0; +x_57 = 1; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_58 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_56, x_57, x_52, x_8, x_9, x_10, x_11, x_12, x_13, x_49); +x_58 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_56, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_49); if (lean_obj_tag(x_58) == 0) { lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; @@ -16878,8 +16878,9 @@ x_66 = lean_ctor_get(x_58, 1); x_67 = l_Lean_Exception_isRuntime(x_65); if (x_67 == 0) { -lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +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_free_object(x_58); +x_68 = 0; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -16887,31 +16888,31 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_48); -x_68 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_66); -x_69 = lean_ctor_get(x_68, 1); -lean_inc(x_69); -lean_dec(x_68); +x_69 = l_Lean_Elab_Term_SavedState_restore(x_48, x_68, x_8, x_9, x_10, x_11, x_12, x_13, x_66); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); lean_inc(x_39); -x_70 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); -lean_closure_set(x_70, 0, x_39); -lean_closure_set(x_70, 1, x_36); -x_71 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); -lean_closure_set(x_71, 0, x_70); +x_71 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); +lean_closure_set(x_71, 0, x_39); +lean_closure_set(x_71, 1, x_36); +x_72 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); +lean_closure_set(x_72, 0, x_71); 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_Elab_Term_commitIfNoErrors_x3f___rarg(x_71, x_8, x_9, x_10, x_11, x_12, x_13, x_69); -if (lean_obj_tag(x_72) == 0) -{ -lean_object* x_73; -x_73 = lean_ctor_get(x_72, 0); -lean_inc(x_73); +x_73 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_72, x_8, x_9, x_10, x_11, x_12, x_13, x_70); if (lean_obj_tag(x_73) == 0) { -uint8_t x_74; +lean_object* x_74; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +if (lean_obj_tag(x_74) == 0) +{ +uint8_t x_75; lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -16922,112 +16923,112 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_74 = !lean_is_exclusive(x_72); -if (x_74 == 0) +x_75 = !lean_is_exclusive(x_73); +if (x_75 == 0) { -lean_object* x_75; -x_75 = lean_ctor_get(x_72, 0); -lean_dec(x_75); -lean_ctor_set_tag(x_72, 1); -lean_ctor_set(x_72, 0, x_65); -return x_72; +lean_object* x_76; +x_76 = lean_ctor_get(x_73, 0); +lean_dec(x_76); +lean_ctor_set_tag(x_73, 1); +lean_ctor_set(x_73, 0, x_65); +return x_73; } else { -lean_object* x_76; lean_object* x_77; -x_76 = lean_ctor_get(x_72, 1); -lean_inc(x_76); -lean_dec(x_72); -x_77 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_77, 0, x_65); -lean_ctor_set(x_77, 1, x_76); -return x_77; +lean_object* x_77; lean_object* x_78; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_78, 0, x_65); +lean_ctor_set(x_78, 1, x_77); +return x_78; } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_72, 1); -lean_inc(x_78); -lean_dec(x_72); -x_79 = lean_ctor_get(x_73, 0); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_73, 1); lean_inc(x_79); lean_dec(x_73); +x_80 = lean_ctor_get(x_74, 0); +lean_inc(x_80); +lean_dec(x_74); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_80 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_79, x_39, x_10, x_11, x_12, x_13, x_78); -if (lean_obj_tag(x_80) == 0) -{ -lean_object* x_81; -x_81 = lean_ctor_get(x_80, 0); -lean_inc(x_81); +x_81 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_80, x_39, x_10, x_11, x_12, x_13, x_79); if (lean_obj_tag(x_81) == 0) { -lean_object* x_82; lean_object* x_83; uint8_t x_84; +lean_object* x_82; +x_82 = lean_ctor_get(x_81, 0); +lean_inc(x_82); +if (lean_obj_tag(x_82) == 0) +{ +lean_object* x_83; lean_object* x_84; uint8_t x_85; lean_dec(x_19); lean_dec(x_4); -x_82 = lean_ctor_get(x_80, 1); -lean_inc(x_82); -lean_dec(x_80); -x_83 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_82); -x_84 = !lean_is_exclusive(x_83); -if (x_84 == 0) +x_83 = lean_ctor_get(x_81, 1); +lean_inc(x_83); +lean_dec(x_81); +x_84 = l_Lean_Elab_Term_SavedState_restore(x_48, x_68, x_8, x_9, x_10, x_11, x_12, x_13, x_83); +x_85 = !lean_is_exclusive(x_84); +if (x_85 == 0) { -lean_object* x_85; -x_85 = lean_ctor_get(x_83, 0); -lean_dec(x_85); -lean_ctor_set_tag(x_83, 1); -lean_ctor_set(x_83, 0, x_65); -return x_83; +lean_object* x_86; +x_86 = lean_ctor_get(x_84, 0); +lean_dec(x_86); +lean_ctor_set_tag(x_84, 1); +lean_ctor_set(x_84, 0, x_65); +return x_84; } else { -lean_object* x_86; lean_object* x_87; -x_86 = lean_ctor_get(x_83, 1); -lean_inc(x_86); -lean_dec(x_83); -x_87 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_87, 0, x_65); -lean_ctor_set(x_87, 1, x_86); -return x_87; +lean_object* x_87; lean_object* x_88; +x_87 = lean_ctor_get(x_84, 1); +lean_inc(x_87); +lean_dec(x_84); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_65); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } 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; -x_88 = lean_ctor_get(x_80, 1); -lean_inc(x_88); -lean_dec(x_80); -x_89 = lean_ctor_get(x_81, 0); +lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; +x_89 = lean_ctor_get(x_81, 1); lean_inc(x_89); lean_dec(x_81); +x_90 = lean_ctor_get(x_82, 0); +lean_inc(x_90); +lean_dec(x_82); 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_90 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_88); -x_91 = lean_ctor_get(x_90, 1); -lean_inc(x_91); -lean_dec(x_90); +x_91 = l_Lean_Elab_Term_SavedState_restore(x_48, x_68, x_8, x_9, x_10, x_11, x_12, x_13, x_89); +x_92 = lean_ctor_get(x_91, 1); +lean_inc(x_92); +lean_dec(x_91); lean_inc(x_4); -x_92 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_92, 0, x_65); -lean_ctor_set(x_92, 1, x_4); -lean_ctor_set(x_92, 2, x_89); -x_93 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_93, 0, x_92); -x_41 = x_93; -x_42 = x_91; +x_93 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_93, 0, x_65); +lean_ctor_set(x_93, 1, x_4); +lean_ctor_set(x_93, 2, x_90); +x_94 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_94, 0, x_93); +x_41 = x_94; +x_42 = x_92; goto block_46; } } else { -uint8_t x_94; +uint8_t x_95; lean_dec(x_65); lean_dec(x_48); lean_dec(x_19); @@ -17038,30 +17039,30 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_94 = !lean_is_exclusive(x_80); -if (x_94 == 0) +x_95 = !lean_is_exclusive(x_81); +if (x_95 == 0) { -return x_80; +return x_81; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_80, 0); -x_96 = lean_ctor_get(x_80, 1); +lean_object* x_96; lean_object* x_97; lean_object* x_98; +x_96 = lean_ctor_get(x_81, 0); +x_97 = lean_ctor_get(x_81, 1); +lean_inc(x_97); lean_inc(x_96); -lean_inc(x_95); -lean_dec(x_80); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_dec(x_81); +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_98; +uint8_t x_99; lean_dec(x_65); lean_dec(x_48); lean_dec(x_39); @@ -17073,31 +17074,31 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_98 = !lean_is_exclusive(x_72); -if (x_98 == 0) +x_99 = !lean_is_exclusive(x_73); +if (x_99 == 0) { -return x_72; +return x_73; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_72, 0); -x_100 = lean_ctor_get(x_72, 1); +lean_object* x_100; lean_object* x_101; lean_object* x_102; +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_inc(x_99); -lean_dec(x_72); -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; +lean_dec(x_73); +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; } } } else { -uint8_t x_102; -x_102 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); -if (x_102 == 0) +uint8_t x_103; +x_103 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); +if (x_103 == 0) { lean_dec(x_48); lean_dec(x_39); @@ -17114,8 +17115,9 @@ return x_58; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +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_free_object(x_58); +x_104 = 0; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -17123,31 +17125,31 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_48); -x_103 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_66); -x_104 = lean_ctor_get(x_103, 1); -lean_inc(x_104); -lean_dec(x_103); +x_105 = l_Lean_Elab_Term_SavedState_restore(x_48, x_104, x_8, x_9, x_10, x_11, x_12, x_13, x_66); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); lean_inc(x_39); -x_105 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); -lean_closure_set(x_105, 0, x_39); -lean_closure_set(x_105, 1, x_36); -x_106 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); -lean_closure_set(x_106, 0, x_105); +x_107 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); +lean_closure_set(x_107, 0, x_39); +lean_closure_set(x_107, 1, x_36); +x_108 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); +lean_closure_set(x_108, 0, x_107); 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_107 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_106, x_8, x_9, x_10, x_11, x_12, x_13, x_104); -if (lean_obj_tag(x_107) == 0) +x_109 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_108, x_8, x_9, x_10, x_11, x_12, x_13, x_106); +if (lean_obj_tag(x_109) == 0) { -lean_object* x_108; -x_108 = lean_ctor_get(x_107, 0); -lean_inc(x_108); -if (lean_obj_tag(x_108) == 0) +lean_object* x_110; +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +if (lean_obj_tag(x_110) == 0) { -uint8_t x_109; +uint8_t x_111; lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -17158,112 +17160,112 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_109 = !lean_is_exclusive(x_107); -if (x_109 == 0) +x_111 = !lean_is_exclusive(x_109); +if (x_111 == 0) { -lean_object* x_110; -x_110 = lean_ctor_get(x_107, 0); -lean_dec(x_110); -lean_ctor_set_tag(x_107, 1); -lean_ctor_set(x_107, 0, x_65); -return x_107; +lean_object* x_112; +x_112 = lean_ctor_get(x_109, 0); +lean_dec(x_112); +lean_ctor_set_tag(x_109, 1); +lean_ctor_set(x_109, 0, x_65); +return x_109; } else { -lean_object* x_111; lean_object* x_112; -x_111 = lean_ctor_get(x_107, 1); -lean_inc(x_111); -lean_dec(x_107); -x_112 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_112, 0, x_65); -lean_ctor_set(x_112, 1, x_111); -return x_112; +lean_object* x_113; lean_object* x_114; +x_113 = lean_ctor_get(x_109, 1); +lean_inc(x_113); +lean_dec(x_109); +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_65); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; -x_113 = lean_ctor_get(x_107, 1); -lean_inc(x_113); -lean_dec(x_107); -x_114 = lean_ctor_get(x_108, 0); -lean_inc(x_114); -lean_dec(x_108); +lean_object* x_115; lean_object* x_116; lean_object* x_117; +x_115 = lean_ctor_get(x_109, 1); +lean_inc(x_115); +lean_dec(x_109); +x_116 = lean_ctor_get(x_110, 0); +lean_inc(x_116); +lean_dec(x_110); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_115 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_114, x_39, x_10, x_11, x_12, x_13, x_113); -if (lean_obj_tag(x_115) == 0) +x_117 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_116, x_39, x_10, x_11, x_12, x_13, x_115); +if (lean_obj_tag(x_117) == 0) { -lean_object* x_116; -x_116 = lean_ctor_get(x_115, 0); -lean_inc(x_116); -if (lean_obj_tag(x_116) == 0) +lean_object* x_118; +x_118 = lean_ctor_get(x_117, 0); +lean_inc(x_118); +if (lean_obj_tag(x_118) == 0) { -lean_object* x_117; lean_object* x_118; uint8_t x_119; +lean_object* x_119; lean_object* x_120; uint8_t x_121; lean_dec(x_19); lean_dec(x_4); -x_117 = lean_ctor_get(x_115, 1); -lean_inc(x_117); -lean_dec(x_115); -x_118 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_117); -x_119 = !lean_is_exclusive(x_118); -if (x_119 == 0) +x_119 = lean_ctor_get(x_117, 1); +lean_inc(x_119); +lean_dec(x_117); +x_120 = l_Lean_Elab_Term_SavedState_restore(x_48, x_104, x_8, x_9, x_10, x_11, x_12, x_13, x_119); +x_121 = !lean_is_exclusive(x_120); +if (x_121 == 0) { -lean_object* x_120; -x_120 = lean_ctor_get(x_118, 0); -lean_dec(x_120); -lean_ctor_set_tag(x_118, 1); -lean_ctor_set(x_118, 0, x_65); -return x_118; +lean_object* x_122; +x_122 = lean_ctor_get(x_120, 0); +lean_dec(x_122); +lean_ctor_set_tag(x_120, 1); +lean_ctor_set(x_120, 0, x_65); +return x_120; } else { -lean_object* x_121; lean_object* x_122; -x_121 = lean_ctor_get(x_118, 1); -lean_inc(x_121); -lean_dec(x_118); -x_122 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_122, 0, x_65); -lean_ctor_set(x_122, 1, x_121); -return x_122; +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_120, 1); +lean_inc(x_123); +lean_dec(x_120); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_65); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } 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; -x_123 = lean_ctor_get(x_115, 1); -lean_inc(x_123); -lean_dec(x_115); -x_124 = lean_ctor_get(x_116, 0); -lean_inc(x_124); -lean_dec(x_116); +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; +x_125 = lean_ctor_get(x_117, 1); +lean_inc(x_125); +lean_dec(x_117); +x_126 = lean_ctor_get(x_118, 0); +lean_inc(x_126); +lean_dec(x_118); 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_125 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_123); -x_126 = lean_ctor_get(x_125, 1); -lean_inc(x_126); -lean_dec(x_125); +x_127 = l_Lean_Elab_Term_SavedState_restore(x_48, x_104, x_8, x_9, x_10, x_11, x_12, x_13, x_125); +x_128 = lean_ctor_get(x_127, 1); +lean_inc(x_128); +lean_dec(x_127); lean_inc(x_4); -x_127 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_127, 0, x_65); -lean_ctor_set(x_127, 1, x_4); -lean_ctor_set(x_127, 2, x_124); -x_128 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_128, 0, x_127); -x_41 = x_128; -x_42 = x_126; +x_129 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_129, 0, x_65); +lean_ctor_set(x_129, 1, x_4); +lean_ctor_set(x_129, 2, x_126); +x_130 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_130, 0, x_129); +x_41 = x_130; +x_42 = x_128; goto block_46; } } else { -uint8_t x_129; +uint8_t x_131; lean_dec(x_65); lean_dec(x_48); lean_dec(x_19); @@ -17274,30 +17276,30 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_129 = !lean_is_exclusive(x_115); -if (x_129 == 0) +x_131 = !lean_is_exclusive(x_117); +if (x_131 == 0) { -return x_115; +return x_117; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_115, 0); -x_131 = lean_ctor_get(x_115, 1); -lean_inc(x_131); -lean_inc(x_130); -lean_dec(x_115); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_132 = lean_ctor_get(x_117, 0); +x_133 = lean_ctor_get(x_117, 1); +lean_inc(x_133); +lean_inc(x_132); +lean_dec(x_117); +x_134 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +return x_134; } } } } else { -uint8_t x_133; +uint8_t x_135; lean_dec(x_65); lean_dec(x_48); lean_dec(x_39); @@ -17309,23 +17311,23 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_133 = !lean_is_exclusive(x_107); -if (x_133 == 0) +x_135 = !lean_is_exclusive(x_109); +if (x_135 == 0) { -return x_107; +return x_109; } else { -lean_object* x_134; lean_object* x_135; lean_object* x_136; -x_134 = lean_ctor_get(x_107, 0); -x_135 = lean_ctor_get(x_107, 1); -lean_inc(x_135); -lean_inc(x_134); -lean_dec(x_107); -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_136; lean_object* x_137; lean_object* x_138; +x_136 = lean_ctor_get(x_109, 0); +x_137 = lean_ctor_get(x_109, 1); +lean_inc(x_137); +lean_inc(x_136); +lean_dec(x_109); +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; } } } @@ -17333,16 +17335,17 @@ return x_136; } else { -lean_object* x_137; lean_object* x_138; uint8_t x_139; -x_137 = lean_ctor_get(x_58, 0); -x_138 = lean_ctor_get(x_58, 1); -lean_inc(x_138); -lean_inc(x_137); +lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_139 = lean_ctor_get(x_58, 0); +x_140 = lean_ctor_get(x_58, 1); +lean_inc(x_140); +lean_inc(x_139); lean_dec(x_58); -x_139 = l_Lean_Exception_isRuntime(x_137); -if (x_139 == 0) +x_141 = l_Lean_Exception_isRuntime(x_139); +if (x_141 == 0) { -lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; +uint8_t x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_142 = 0; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -17350,31 +17353,31 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_48); -x_140 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_138); -x_141 = lean_ctor_get(x_140, 1); -lean_inc(x_141); -lean_dec(x_140); +x_143 = l_Lean_Elab_Term_SavedState_restore(x_48, x_142, x_8, x_9, x_10, x_11, x_12, x_13, x_140); +x_144 = lean_ctor_get(x_143, 1); +lean_inc(x_144); +lean_dec(x_143); lean_inc(x_39); -x_142 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); -lean_closure_set(x_142, 0, x_39); -lean_closure_set(x_142, 1, x_36); -x_143 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); -lean_closure_set(x_143, 0, x_142); +x_145 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); +lean_closure_set(x_145, 0, x_39); +lean_closure_set(x_145, 1, x_36); +x_146 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); +lean_closure_set(x_146, 0, x_145); 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_144 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_143, x_8, x_9, x_10, x_11, x_12, x_13, x_141); -if (lean_obj_tag(x_144) == 0) +x_147 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_146, x_8, x_9, x_10, x_11, x_12, x_13, x_144); +if (lean_obj_tag(x_147) == 0) { -lean_object* x_145; -x_145 = lean_ctor_get(x_144, 0); -lean_inc(x_145); -if (lean_obj_tag(x_145) == 0) +lean_object* x_148; +x_148 = lean_ctor_get(x_147, 0); +lean_inc(x_148); +if (lean_obj_tag(x_148) == 0) { -lean_object* x_146; lean_object* x_147; lean_object* x_148; +lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -17385,109 +17388,109 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_146 = lean_ctor_get(x_144, 1); -lean_inc(x_146); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_147 = x_144; +x_149 = lean_ctor_get(x_147, 1); +lean_inc(x_149); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_150 = x_147; } else { - lean_dec_ref(x_144); - x_147 = lean_box(0); + lean_dec_ref(x_147); + x_150 = lean_box(0); } -if (lean_is_scalar(x_147)) { - x_148 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_150)) { + x_151 = lean_alloc_ctor(1, 2, 0); } else { - x_148 = x_147; - lean_ctor_set_tag(x_148, 1); + x_151 = x_150; + lean_ctor_set_tag(x_151, 1); } -lean_ctor_set(x_148, 0, x_137); -lean_ctor_set(x_148, 1, x_146); -return x_148; +lean_ctor_set(x_151, 0, x_139); +lean_ctor_set(x_151, 1, x_149); +return x_151; } else { -lean_object* x_149; lean_object* x_150; lean_object* x_151; -x_149 = lean_ctor_get(x_144, 1); -lean_inc(x_149); -lean_dec(x_144); -x_150 = lean_ctor_get(x_145, 0); -lean_inc(x_150); -lean_dec(x_145); +lean_object* x_152; lean_object* x_153; lean_object* x_154; +x_152 = lean_ctor_get(x_147, 1); +lean_inc(x_152); +lean_dec(x_147); +x_153 = lean_ctor_get(x_148, 0); +lean_inc(x_153); +lean_dec(x_148); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_151 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_150, x_39, x_10, x_11, x_12, x_13, x_149); -if (lean_obj_tag(x_151) == 0) +x_154 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_153, x_39, x_10, x_11, x_12, x_13, x_152); +if (lean_obj_tag(x_154) == 0) { -lean_object* x_152; -x_152 = lean_ctor_get(x_151, 0); -lean_inc(x_152); -if (lean_obj_tag(x_152) == 0) +lean_object* x_155; +x_155 = lean_ctor_get(x_154, 0); +lean_inc(x_155); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; +lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_dec(x_19); lean_dec(x_4); -x_153 = lean_ctor_get(x_151, 1); -lean_inc(x_153); -lean_dec(x_151); -x_154 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_153); -x_155 = lean_ctor_get(x_154, 1); -lean_inc(x_155); -if (lean_is_exclusive(x_154)) { - lean_ctor_release(x_154, 0); - lean_ctor_release(x_154, 1); - x_156 = x_154; +x_156 = lean_ctor_get(x_154, 1); +lean_inc(x_156); +lean_dec(x_154); +x_157 = l_Lean_Elab_Term_SavedState_restore(x_48, x_142, x_8, x_9, x_10, x_11, x_12, x_13, x_156); +x_158 = lean_ctor_get(x_157, 1); +lean_inc(x_158); +if (lean_is_exclusive(x_157)) { + lean_ctor_release(x_157, 0); + lean_ctor_release(x_157, 1); + x_159 = x_157; } else { - lean_dec_ref(x_154); - x_156 = lean_box(0); + lean_dec_ref(x_157); + x_159 = lean_box(0); } -if (lean_is_scalar(x_156)) { - x_157 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_159)) { + x_160 = lean_alloc_ctor(1, 2, 0); } else { - x_157 = x_156; - lean_ctor_set_tag(x_157, 1); + x_160 = x_159; + lean_ctor_set_tag(x_160, 1); } -lean_ctor_set(x_157, 0, x_137); -lean_ctor_set(x_157, 1, x_155); -return x_157; +lean_ctor_set(x_160, 0, x_139); +lean_ctor_set(x_160, 1, x_158); +return x_160; } else { -lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_158 = lean_ctor_get(x_151, 1); -lean_inc(x_158); -lean_dec(x_151); -x_159 = lean_ctor_get(x_152, 0); -lean_inc(x_159); -lean_dec(x_152); +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; +x_161 = lean_ctor_get(x_154, 1); +lean_inc(x_161); +lean_dec(x_154); +x_162 = lean_ctor_get(x_155, 0); +lean_inc(x_162); +lean_dec(x_155); 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_160 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_158); -x_161 = lean_ctor_get(x_160, 1); -lean_inc(x_161); -lean_dec(x_160); +x_163 = l_Lean_Elab_Term_SavedState_restore(x_48, x_142, x_8, x_9, x_10, x_11, x_12, x_13, x_161); +x_164 = lean_ctor_get(x_163, 1); +lean_inc(x_164); +lean_dec(x_163); lean_inc(x_4); -x_162 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_162, 0, x_137); -lean_ctor_set(x_162, 1, x_4); -lean_ctor_set(x_162, 2, x_159); -x_163 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_163, 0, x_162); -x_41 = x_163; -x_42 = x_161; +x_165 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_165, 0, x_139); +lean_ctor_set(x_165, 1, x_4); +lean_ctor_set(x_165, 2, x_162); +x_166 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_166, 0, x_165); +x_41 = x_166; +x_42 = x_164; goto block_46; } } else { -lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; -lean_dec(x_137); +lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; +lean_dec(x_139); lean_dec(x_48); lean_dec(x_19); lean_dec(x_13); @@ -17497,33 +17500,33 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_164 = lean_ctor_get(x_151, 0); -lean_inc(x_164); -x_165 = lean_ctor_get(x_151, 1); -lean_inc(x_165); -if (lean_is_exclusive(x_151)) { - lean_ctor_release(x_151, 0); - lean_ctor_release(x_151, 1); - x_166 = x_151; +x_167 = lean_ctor_get(x_154, 0); +lean_inc(x_167); +x_168 = lean_ctor_get(x_154, 1); +lean_inc(x_168); +if (lean_is_exclusive(x_154)) { + lean_ctor_release(x_154, 0); + lean_ctor_release(x_154, 1); + x_169 = x_154; } else { - lean_dec_ref(x_151); - x_166 = lean_box(0); + lean_dec_ref(x_154); + x_169 = lean_box(0); } -if (lean_is_scalar(x_166)) { - x_167 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_169)) { + x_170 = lean_alloc_ctor(1, 2, 0); } else { - x_167 = x_166; + x_170 = x_169; } -lean_ctor_set(x_167, 0, x_164); -lean_ctor_set(x_167, 1, x_165); -return x_167; +lean_ctor_set(x_170, 0, x_167); +lean_ctor_set(x_170, 1, x_168); +return x_170; } } } else { -lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; -lean_dec(x_137); +lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; +lean_dec(x_139); lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -17534,35 +17537,35 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_168 = lean_ctor_get(x_144, 0); -lean_inc(x_168); -x_169 = lean_ctor_get(x_144, 1); -lean_inc(x_169); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_170 = x_144; +x_171 = lean_ctor_get(x_147, 0); +lean_inc(x_171); +x_172 = lean_ctor_get(x_147, 1); +lean_inc(x_172); +if (lean_is_exclusive(x_147)) { + lean_ctor_release(x_147, 0); + lean_ctor_release(x_147, 1); + x_173 = x_147; } else { - lean_dec_ref(x_144); - x_170 = lean_box(0); + lean_dec_ref(x_147); + x_173 = lean_box(0); } -if (lean_is_scalar(x_170)) { - x_171 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_173)) { + x_174 = lean_alloc_ctor(1, 2, 0); } else { - x_171 = x_170; + x_174 = x_173; } -lean_ctor_set(x_171, 0, x_168); -lean_ctor_set(x_171, 1, x_169); -return x_171; +lean_ctor_set(x_174, 0, x_171); +lean_ctor_set(x_174, 1, x_172); +return x_174; } } else { -uint8_t x_172; -x_172 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); -if (x_172 == 0) +uint8_t x_175; +x_175 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); +if (x_175 == 0) { -lean_object* x_173; +lean_object* x_176; lean_dec(x_48); lean_dec(x_39); lean_dec(x_36); @@ -17574,14 +17577,15 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_173 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_173, 0, x_137); -lean_ctor_set(x_173, 1, x_138); -return x_173; +x_176 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_176, 0, x_139); +lean_ctor_set(x_176, 1, x_140); +return x_176; } else { -lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; +uint8_t x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_177 = 0; lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); @@ -17589,31 +17593,31 @@ lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_48); -x_174 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_138); -x_175 = lean_ctor_get(x_174, 1); -lean_inc(x_175); -lean_dec(x_174); +x_178 = l_Lean_Elab_Term_SavedState_restore(x_48, x_177, x_8, x_9, x_10, x_11, x_12, x_13, x_140); +x_179 = lean_ctor_get(x_178, 1); +lean_inc(x_179); +lean_dec(x_178); lean_inc(x_39); -x_176 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); -lean_closure_set(x_176, 0, x_39); -lean_closure_set(x_176, 1, x_36); -x_177 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); -lean_closure_set(x_177, 0, x_176); +x_180 = lean_alloc_closure((void*)(l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___lambda__2), 9, 2); +lean_closure_set(x_180, 0, x_39); +lean_closure_set(x_180, 1, x_36); +x_181 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_withPatternElabConfig___rarg), 8, 1); +lean_closure_set(x_181, 0, x_180); 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_178 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_177, x_8, x_9, x_10, x_11, x_12, x_13, x_175); -if (lean_obj_tag(x_178) == 0) +x_182 = l_Lean_Elab_Term_commitIfNoErrors_x3f___rarg(x_181, x_8, x_9, x_10, x_11, x_12, x_13, x_179); +if (lean_obj_tag(x_182) == 0) { -lean_object* x_179; -x_179 = lean_ctor_get(x_178, 0); -lean_inc(x_179); -if (lean_obj_tag(x_179) == 0) +lean_object* x_183; +x_183 = lean_ctor_get(x_182, 0); +lean_inc(x_183); +if (lean_obj_tag(x_183) == 0) { -lean_object* x_180; lean_object* x_181; lean_object* x_182; +lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -17624,109 +17628,109 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_180 = lean_ctor_get(x_178, 1); -lean_inc(x_180); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_181 = x_178; +x_184 = lean_ctor_get(x_182, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_185 = x_182; } else { - lean_dec_ref(x_178); - x_181 = lean_box(0); + lean_dec_ref(x_182); + x_185 = lean_box(0); } -if (lean_is_scalar(x_181)) { - x_182 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); } else { - x_182 = x_181; - lean_ctor_set_tag(x_182, 1); + x_186 = x_185; + lean_ctor_set_tag(x_186, 1); } -lean_ctor_set(x_182, 0, x_137); -lean_ctor_set(x_182, 1, x_180); -return x_182; +lean_ctor_set(x_186, 0, x_139); +lean_ctor_set(x_186, 1, x_184); +return x_186; } else { -lean_object* x_183; lean_object* x_184; lean_object* x_185; -x_183 = lean_ctor_get(x_178, 1); -lean_inc(x_183); -lean_dec(x_178); -x_184 = lean_ctor_get(x_179, 0); -lean_inc(x_184); -lean_dec(x_179); +lean_object* x_187; lean_object* x_188; lean_object* x_189; +x_187 = lean_ctor_get(x_182, 1); +lean_inc(x_187); +lean_dec(x_182); +x_188 = lean_ctor_get(x_183, 0); +lean_inc(x_188); +lean_dec(x_183); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_185 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_184, x_39, x_10, x_11, x_12, x_13, x_183); -if (lean_obj_tag(x_185) == 0) +x_189 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_findDiscrRefinementPath(x_188, x_39, x_10, x_11, x_12, x_13, x_187); +if (lean_obj_tag(x_189) == 0) { -lean_object* x_186; -x_186 = lean_ctor_get(x_185, 0); -lean_inc(x_186); -if (lean_obj_tag(x_186) == 0) +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_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_dec(x_19); lean_dec(x_4); -x_187 = lean_ctor_get(x_185, 1); -lean_inc(x_187); -lean_dec(x_185); -x_188 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_187); -x_189 = lean_ctor_get(x_188, 1); -lean_inc(x_189); -if (lean_is_exclusive(x_188)) { - lean_ctor_release(x_188, 0); - lean_ctor_release(x_188, 1); - x_190 = x_188; +x_191 = lean_ctor_get(x_189, 1); +lean_inc(x_191); +lean_dec(x_189); +x_192 = l_Lean_Elab_Term_SavedState_restore(x_48, x_177, x_8, x_9, x_10, x_11, x_12, x_13, x_191); +x_193 = lean_ctor_get(x_192, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_192)) { + lean_ctor_release(x_192, 0); + lean_ctor_release(x_192, 1); + x_194 = x_192; } else { - lean_dec_ref(x_188); - x_190 = lean_box(0); + lean_dec_ref(x_192); + x_194 = lean_box(0); } -if (lean_is_scalar(x_190)) { - x_191 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); } else { - x_191 = x_190; - lean_ctor_set_tag(x_191, 1); + x_195 = x_194; + lean_ctor_set_tag(x_195, 1); } -lean_ctor_set(x_191, 0, x_137); -lean_ctor_set(x_191, 1, x_189); -return x_191; +lean_ctor_set(x_195, 0, x_139); +lean_ctor_set(x_195, 1, x_193); +return x_195; } else { -lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; -x_192 = lean_ctor_get(x_185, 1); -lean_inc(x_192); -lean_dec(x_185); -x_193 = lean_ctor_get(x_186, 0); -lean_inc(x_193); -lean_dec(x_186); +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_196 = lean_ctor_get(x_189, 1); +lean_inc(x_196); +lean_dec(x_189); +x_197 = lean_ctor_get(x_190, 0); +lean_inc(x_197); +lean_dec(x_190); 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_194 = l_Lean_Elab_Term_SavedState_restore(x_48, x_57, x_8, x_9, x_10, x_11, x_12, x_13, x_192); -x_195 = lean_ctor_get(x_194, 1); -lean_inc(x_195); -lean_dec(x_194); +x_198 = l_Lean_Elab_Term_SavedState_restore(x_48, x_177, x_8, x_9, x_10, x_11, x_12, x_13, x_196); +x_199 = lean_ctor_get(x_198, 1); +lean_inc(x_199); +lean_dec(x_198); lean_inc(x_4); -x_196 = lean_alloc_ctor(0, 3, 0); -lean_ctor_set(x_196, 0, x_137); -lean_ctor_set(x_196, 1, x_4); -lean_ctor_set(x_196, 2, x_193); -x_197 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_197, 0, x_196); -x_41 = x_197; -x_42 = x_195; +x_200 = lean_alloc_ctor(0, 3, 0); +lean_ctor_set(x_200, 0, x_139); +lean_ctor_set(x_200, 1, x_4); +lean_ctor_set(x_200, 2, x_197); +x_201 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_201, 0, x_200); +x_41 = x_201; +x_42 = x_199; goto block_46; } } else { -lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; -lean_dec(x_137); +lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; +lean_dec(x_139); lean_dec(x_48); lean_dec(x_19); lean_dec(x_13); @@ -17736,33 +17740,33 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_198 = lean_ctor_get(x_185, 0); -lean_inc(x_198); -x_199 = lean_ctor_get(x_185, 1); -lean_inc(x_199); -if (lean_is_exclusive(x_185)) { - lean_ctor_release(x_185, 0); - lean_ctor_release(x_185, 1); - x_200 = x_185; +x_202 = lean_ctor_get(x_189, 0); +lean_inc(x_202); +x_203 = lean_ctor_get(x_189, 1); +lean_inc(x_203); +if (lean_is_exclusive(x_189)) { + lean_ctor_release(x_189, 0); + lean_ctor_release(x_189, 1); + x_204 = x_189; } else { - lean_dec_ref(x_185); - x_200 = lean_box(0); + lean_dec_ref(x_189); + x_204 = lean_box(0); } -if (lean_is_scalar(x_200)) { - x_201 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_204)) { + x_205 = lean_alloc_ctor(1, 2, 0); } else { - x_201 = x_200; + x_205 = x_204; } -lean_ctor_set(x_201, 0, x_198); -lean_ctor_set(x_201, 1, x_199); -return x_201; +lean_ctor_set(x_205, 0, x_202); +lean_ctor_set(x_205, 1, x_203); +return x_205; } } } else { -lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; -lean_dec(x_137); +lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; +lean_dec(x_139); lean_dec(x_48); lean_dec(x_39); lean_dec(x_19); @@ -17773,26 +17777,26 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_202 = lean_ctor_get(x_178, 0); -lean_inc(x_202); -x_203 = lean_ctor_get(x_178, 1); -lean_inc(x_203); -if (lean_is_exclusive(x_178)) { - lean_ctor_release(x_178, 0); - lean_ctor_release(x_178, 1); - x_204 = x_178; +x_206 = lean_ctor_get(x_182, 0); +lean_inc(x_206); +x_207 = lean_ctor_get(x_182, 1); +lean_inc(x_207); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_208 = x_182; } else { - lean_dec_ref(x_178); - x_204 = lean_box(0); + lean_dec_ref(x_182); + x_208 = lean_box(0); } -if (lean_is_scalar(x_204)) { - x_205 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_208)) { + x_209 = lean_alloc_ctor(1, 2, 0); } else { - x_205 = x_204; + x_209 = x_208; } -lean_ctor_set(x_205, 0, x_202); -lean_ctor_set(x_205, 1, x_203); -return x_205; +lean_ctor_set(x_209, 0, x_206); +lean_ctor_set(x_209, 1, x_207); +return x_209; } } } @@ -17824,46 +17828,46 @@ goto block_31; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; +lean_object* x_210; lean_object* x_211; lean_object* x_212; uint8_t x_213; lean_dec(x_37); lean_dec(x_36); lean_dec(x_33); lean_dec(x_19); lean_dec(x_4); -x_206 = lean_ctor_get(x_35, 1); -lean_inc(x_206); +x_210 = lean_ctor_get(x_35, 1); +lean_inc(x_210); lean_dec(x_35); -x_207 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___closed__2; -x_208 = l_Lean_throwError___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__1(x_207, x_8, x_9, x_10, x_11, x_12, x_13, x_206); +x_211 = l_Std_Range_forIn_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__2___closed__2; +x_212 = l_Lean_throwError___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___spec__1(x_211, x_8, x_9, x_10, x_11, x_12, x_13, x_210); 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_209 = !lean_is_exclusive(x_208); -if (x_209 == 0) +x_213 = !lean_is_exclusive(x_212); +if (x_213 == 0) { -return x_208; +return x_212; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_210 = lean_ctor_get(x_208, 0); -x_211 = lean_ctor_get(x_208, 1); -lean_inc(x_211); -lean_inc(x_210); -lean_dec(x_208); -x_212 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -return x_212; +lean_object* x_214; lean_object* x_215; lean_object* x_216; +x_214 = lean_ctor_get(x_212, 0); +x_215 = lean_ctor_get(x_212, 1); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_212); +x_216 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_216, 0, x_214); +lean_ctor_set(x_216, 1, x_215); +return x_216; } } } else { -uint8_t x_213; +uint8_t x_217; lean_dec(x_36); lean_dec(x_33); lean_dec(x_19); @@ -17874,30 +17878,30 @@ lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); -x_213 = !lean_is_exclusive(x_35); -if (x_213 == 0) +x_217 = !lean_is_exclusive(x_35); +if (x_217 == 0) { return x_35; } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_214 = lean_ctor_get(x_35, 0); -x_215 = lean_ctor_get(x_35, 1); -lean_inc(x_215); -lean_inc(x_214); +lean_object* x_218; lean_object* x_219; lean_object* x_220; +x_218 = lean_ctor_get(x_35, 0); +x_219 = lean_ctor_get(x_35, 1); +lean_inc(x_219); +lean_inc(x_218); lean_dec(x_35); -x_216 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_216, 0, x_214); -lean_ctor_set(x_216, 1, x_215); -return x_216; +x_220 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_220, 0, x_218); +lean_ctor_set(x_220, 1, x_219); +return x_220; } } } } else { -lean_object* x_221; lean_object* x_222; +lean_object* x_225; lean_object* x_226; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -17906,17 +17910,17 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_221 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_221, 0, x_7); -x_222 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_222, 0, x_221); -lean_ctor_set(x_222, 1, x_14); -return x_222; +x_225 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_225, 0, x_7); +x_226 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_226, 0, x_225); +lean_ctor_set(x_226, 1, x_14); +return x_226; } } else { -lean_object* x_223; lean_object* x_224; +lean_object* x_227; lean_object* x_228; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -17925,12 +17929,12 @@ lean_dec(x_9); lean_dec(x_8); lean_dec(x_4); lean_dec(x_3); -x_223 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_223, 0, x_7); -x_224 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_224, 0, x_223); -lean_ctor_set(x_224, 1, x_14); -return x_224; +x_227 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_227, 0, x_7); +x_228 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_228, 0, x_227); +lean_ctor_set(x_228, 1, x_14); +return x_228; } } } @@ -31086,27 +31090,26 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___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) { _start: { -lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; uint8_t x_14; lean_object* x_15; x_13 = lean_alloc_closure((void*)(l___private_Lean_Elab_Match_0__Lean_Elab_Term_elabPatterns___boxed), 9, 2); lean_closure_set(x_13, 0, x_3); lean_closure_set(x_13, 1, x_4); -x_14 = 0; -x_15 = 1; +x_14 = 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 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, x_14, x_15, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -if (lean_obj_tag(x_16) == 0) +x_15 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, 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; -x_17 = lean_ctor_get(x_16, 0); -lean_inc(x_17); -if (lean_obj_tag(x_17) == 0) +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) { -uint8_t x_18; +uint8_t x_17; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -31116,373 +31119,373 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_18 = !lean_is_exclusive(x_16); -if (x_18 == 0) +x_17 = !lean_is_exclusive(x_15); +if (x_17 == 0) { -lean_object* x_19; uint8_t x_20; -x_19 = lean_ctor_get(x_16, 0); -lean_dec(x_19); -x_20 = !lean_is_exclusive(x_17); -if (x_20 == 0) +lean_object* x_18; uint8_t x_19; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +x_19 = !lean_is_exclusive(x_16); +if (x_19 == 0) { -return x_16; +return x_15; } else { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_17, 0); -lean_inc(x_21); -lean_dec(x_17); -x_22 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_16, 0, x_22); -return x_16; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_16, 0); +lean_inc(x_20); +lean_dec(x_16); +x_21 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_21, 0, x_20); +lean_ctor_set(x_15, 0, x_21); +return x_15; } } else { -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_16, 1); +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_15, 1); +lean_inc(x_22); +lean_dec(x_15); +x_23 = lean_ctor_get(x_16, 0); lean_inc(x_23); -lean_dec(x_16); -x_24 = lean_ctor_get(x_17, 0); -lean_inc(x_24); -if (lean_is_exclusive(x_17)) { - lean_ctor_release(x_17, 0); - x_25 = x_17; +if (lean_is_exclusive(x_16)) { + lean_ctor_release(x_16, 0); + x_24 = x_16; } else { - lean_dec_ref(x_17); - x_25 = lean_box(0); + lean_dec_ref(x_16); + x_24 = lean_box(0); } -if (lean_is_scalar(x_25)) { - x_26 = lean_alloc_ctor(0, 1, 0); +if (lean_is_scalar(x_24)) { + x_25 = lean_alloc_ctor(0, 1, 0); } else { - x_26 = x_25; + x_25 = x_24; } -lean_ctor_set(x_26, 0, x_24); -x_27 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_23); -return x_27; +lean_ctor_set(x_25, 0, x_23); +x_26 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_26, 0, x_25); +lean_ctor_set(x_26, 1, x_22); +return x_26; } } else { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_17); -if (x_28 == 0) +uint8_t x_27; +x_27 = !lean_is_exclusive(x_16); +if (x_27 == 0) { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_29 = lean_ctor_get(x_17, 0); -x_30 = lean_ctor_get(x_16, 1); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_28 = lean_ctor_get(x_16, 0); +x_29 = lean_ctor_get(x_15, 1); +lean_inc(x_29); +lean_dec(x_15); +x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); -lean_dec(x_16); -x_31 = lean_ctor_get(x_29, 0); +x_31 = lean_ctor_get(x_28, 1); lean_inc(x_31); -x_32 = lean_ctor_get(x_29, 1); -lean_inc(x_32); -lean_dec(x_29); -x_33 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; -x_34 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_33, x_6, x_7, x_8, x_9, x_10, x_11, x_30); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_unbox(x_35); -lean_dec(x_35); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); +lean_dec(x_28); +x_32 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; +x_33 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_32, x_6, x_7, x_8, x_9, x_10, x_11, x_29); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); lean_dec(x_34); -x_38 = lean_box(0); -x_39 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_31, x_32, x_38, x_6, x_7, x_8, x_9, x_10, x_11, x_37); -if (lean_obj_tag(x_39) == 0) +if (x_35 == 0) { -uint8_t x_40; -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 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___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_30, x_31, x_37, x_6, x_7, x_8, x_9, x_10, x_11, x_36); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_41; -x_41 = lean_ctor_get(x_39, 0); -lean_ctor_set(x_17, 0, x_41); -lean_ctor_set(x_39, 0, x_17); -return x_39; +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) +{ +lean_object* x_40; +x_40 = lean_ctor_get(x_38, 0); +lean_ctor_set(x_16, 0, x_40); +lean_ctor_set(x_38, 0, x_16); +return x_38; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_39, 0); -x_43 = lean_ctor_get(x_39, 1); -lean_inc(x_43); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_38, 0); +x_42 = lean_ctor_get(x_38, 1); lean_inc(x_42); -lean_dec(x_39); -lean_ctor_set(x_17, 0, x_42); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_17); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_inc(x_41); +lean_dec(x_38); +lean_ctor_set(x_16, 0, x_41); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_16); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } else { -uint8_t x_45; -lean_free_object(x_17); -x_45 = !lean_is_exclusive(x_39); -if (x_45 == 0) +uint8_t x_44; +lean_free_object(x_16); +x_44 = !lean_is_exclusive(x_38); +if (x_44 == 0) { -return x_39; +return x_38; } 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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_38, 0); +x_46 = lean_ctor_get(x_38, 1); 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; +lean_inc(x_45); +lean_dec(x_38); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; 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_49 = lean_ctor_get(x_34, 1); -lean_inc(x_49); -lean_dec(x_34); -lean_inc(x_31); -x_50 = lean_array_to_list(lean_box(0), x_31); -x_51 = lean_box(0); -x_52 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_50, x_51); -x_53 = l_Lean_MessageData_ofList(x_52); -lean_dec(x_52); -x_54 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___closed__2; -x_55 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_55, 0, x_54); -lean_ctor_set(x_55, 1, x_53); -x_56 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; -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_Term_traceAtCmdPos___spec__2(x_33, x_57, x_6, x_7, x_8, x_9, x_10, x_11, x_49); -x_59 = lean_ctor_get(x_58, 0); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_48 = lean_ctor_get(x_33, 1); +lean_inc(x_48); +lean_dec(x_33); +lean_inc(x_30); +x_49 = lean_array_to_list(lean_box(0), x_30); +x_50 = lean_box(0); +x_51 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_49, x_50); +x_52 = l_Lean_MessageData_ofList(x_51); +lean_dec(x_51); +x_53 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___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_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; +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_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_32, x_56, x_6, x_7, x_8, x_9, x_10, x_11, x_48); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); lean_inc(x_59); -x_60 = lean_ctor_get(x_58, 1); -lean_inc(x_60); +lean_dec(x_57); +x_60 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_30, x_31, x_58, x_6, x_7, x_8, x_9, x_10, x_11, x_59); lean_dec(x_58); -x_61 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_31, x_32, x_59, x_6, x_7, x_8, x_9, x_10, x_11, x_60); -lean_dec(x_59); -if (lean_obj_tag(x_61) == 0) +if (lean_obj_tag(x_60) == 0) { -uint8_t x_62; -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) +uint8_t x_61; +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) { -lean_object* x_63; -x_63 = lean_ctor_get(x_61, 0); -lean_ctor_set(x_17, 0, x_63); -lean_ctor_set(x_61, 0, x_17); -return x_61; +lean_object* x_62; +x_62 = lean_ctor_get(x_60, 0); +lean_ctor_set(x_16, 0, x_62); +lean_ctor_set(x_60, 0, x_16); +return x_60; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_61, 0); -x_65 = lean_ctor_get(x_61, 1); -lean_inc(x_65); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_60, 0); +x_64 = lean_ctor_get(x_60, 1); lean_inc(x_64); -lean_dec(x_61); -lean_ctor_set(x_17, 0, x_64); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_17); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_inc(x_63); +lean_dec(x_60); +lean_ctor_set(x_16, 0, x_63); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_16); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } else { -uint8_t x_67; -lean_free_object(x_17); -x_67 = !lean_is_exclusive(x_61); -if (x_67 == 0) +uint8_t x_66; +lean_free_object(x_16); +x_66 = !lean_is_exclusive(x_60); +if (x_66 == 0) { -return x_61; +return x_60; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_61, 0); -x_69 = lean_ctor_get(x_61, 1); -lean_inc(x_69); +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_dec(x_61); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_inc(x_67); +lean_dec(x_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; } } } } else { -lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; -x_71 = lean_ctor_get(x_17, 0); +lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; uint8_t x_77; +x_70 = lean_ctor_get(x_16, 0); +lean_inc(x_70); +lean_dec(x_16); +x_71 = lean_ctor_get(x_15, 1); lean_inc(x_71); -lean_dec(x_17); -x_72 = lean_ctor_get(x_16, 1); +lean_dec(x_15); +x_72 = lean_ctor_get(x_70, 0); lean_inc(x_72); -lean_dec(x_16); -x_73 = lean_ctor_get(x_71, 0); +x_73 = lean_ctor_get(x_70, 1); lean_inc(x_73); -x_74 = lean_ctor_get(x_71, 1); -lean_inc(x_74); -lean_dec(x_71); -x_75 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; -x_76 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_75, x_6, x_7, x_8, x_9, x_10, x_11, x_72); -x_77 = lean_ctor_get(x_76, 0); -lean_inc(x_77); -x_78 = lean_unbox(x_77); -lean_dec(x_77); -if (x_78 == 0) -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_76, 1); -lean_inc(x_79); +lean_dec(x_70); +x_74 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__7; +x_75 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_74, x_6, x_7, x_8, x_9, x_10, x_11, x_71); +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_unbox(x_76); lean_dec(x_76); -x_80 = lean_box(0); -x_81 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_73, x_74, x_80, x_6, x_7, x_8, x_9, x_10, x_11, x_79); -if (lean_obj_tag(x_81) == 0) +if (x_77 == 0) { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; -x_82 = lean_ctor_get(x_81, 0); +lean_object* x_78; lean_object* x_79; lean_object* x_80; +x_78 = lean_ctor_get(x_75, 1); +lean_inc(x_78); +lean_dec(x_75); +x_79 = lean_box(0); +x_80 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_72, x_73, x_79, x_6, x_7, x_8, x_9, x_10, x_11, x_78); +if (lean_obj_tag(x_80) == 0) +{ +lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_81 = lean_ctor_get(x_80, 0); +lean_inc(x_81); +x_82 = lean_ctor_get(x_80, 1); lean_inc(x_82); -x_83 = lean_ctor_get(x_81, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_84 = x_81; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_83 = x_80; } else { - lean_dec_ref(x_81); - x_84 = lean_box(0); + lean_dec_ref(x_80); + x_83 = lean_box(0); } -x_85 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_85, 0, x_82); -if (lean_is_scalar(x_84)) { - x_86 = lean_alloc_ctor(0, 2, 0); +x_84 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_84, 0, x_81); +if (lean_is_scalar(x_83)) { + x_85 = lean_alloc_ctor(0, 2, 0); } else { - x_86 = x_84; + x_85 = x_83; } -lean_ctor_set(x_86, 0, x_85); -lean_ctor_set(x_86, 1, x_83); -return x_86; +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_82); +return x_85; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_87 = lean_ctor_get(x_81, 0); +lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_86 = lean_ctor_get(x_80, 0); +lean_inc(x_86); +x_87 = lean_ctor_get(x_80, 1); lean_inc(x_87); -x_88 = lean_ctor_get(x_81, 1); -lean_inc(x_88); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_89 = x_81; +if (lean_is_exclusive(x_80)) { + lean_ctor_release(x_80, 0); + lean_ctor_release(x_80, 1); + x_88 = x_80; } else { - lean_dec_ref(x_81); - x_89 = lean_box(0); + lean_dec_ref(x_80); + x_88 = lean_box(0); } -if (lean_is_scalar(x_89)) { - x_90 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_88)) { + x_89 = lean_alloc_ctor(1, 2, 0); } else { - x_90 = x_89; + x_89 = x_88; } -lean_ctor_set(x_90, 0, x_87); -lean_ctor_set(x_90, 1, x_88); -return x_90; +lean_ctor_set(x_89, 0, x_86); +lean_ctor_set(x_89, 1, x_87); +return x_89; } } else { -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_91 = lean_ctor_get(x_76, 1); -lean_inc(x_91); -lean_dec(x_76); -lean_inc(x_73); -x_92 = lean_array_to_list(lean_box(0), x_73); -x_93 = lean_box(0); -x_94 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_92, x_93); -x_95 = l_Lean_MessageData_ofList(x_94); -lean_dec(x_94); -x_96 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___closed__2; -x_97 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_95); -x_98 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; -x_99 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_99, 0, x_97); -lean_ctor_set(x_99, 1, x_98); -x_100 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_75, x_99, x_6, x_7, x_8, x_9, x_10, x_11, x_91); -x_101 = lean_ctor_get(x_100, 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; lean_object* x_100; lean_object* x_101; lean_object* x_102; +x_90 = lean_ctor_get(x_75, 1); +lean_inc(x_90); +lean_dec(x_75); +lean_inc(x_72); +x_91 = lean_array_to_list(lean_box(0), x_72); +x_92 = lean_box(0); +x_93 = l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(x_91, x_92); +x_94 = l_Lean_MessageData_ofList(x_93); +lean_dec(x_93); +x_95 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___closed__2; +x_96 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_94); +x_97 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_Match_0__Lean_Elab_Term_elabMatchTypeAndDiscrs_elabDiscrsWitMatchType___spec__1___closed__15; +x_98 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_98, 0, x_96); +lean_ctor_set(x_98, 1, x_97); +x_99 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_74, x_98, x_6, x_7, x_8, x_9, x_10, x_11, x_90); +x_100 = lean_ctor_get(x_99, 0); +lean_inc(x_100); +x_101 = lean_ctor_get(x_99, 1); lean_inc(x_101); -x_102 = lean_ctor_get(x_100, 1); -lean_inc(x_102); +lean_dec(x_99); +x_102 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_72, x_73, x_100, x_6, x_7, x_8, x_9, x_10, x_11, x_101); lean_dec(x_100); -x_103 = l___private_Lean_Elab_Match_0__Lean_Elab_Term_withElaboratedLHS___rarg___lambda__2(x_1, x_5, x_2, x_73, x_74, x_101, x_6, x_7, x_8, x_9, x_10, x_11, x_102); -lean_dec(x_101); -if (lean_obj_tag(x_103) == 0) +if (lean_obj_tag(x_102) == 0) { -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_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_103 = lean_ctor_get(x_102, 0); +lean_inc(x_103); +x_104 = lean_ctor_get(x_102, 1); 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; +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_105 = x_102; } else { - lean_dec_ref(x_103); - x_106 = lean_box(0); + lean_dec_ref(x_102); + x_105 = lean_box(0); } -x_107 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_107, 0, x_104); -if (lean_is_scalar(x_106)) { - x_108 = lean_alloc_ctor(0, 2, 0); +x_106 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_106, 0, x_103); +if (lean_is_scalar(x_105)) { + x_107 = lean_alloc_ctor(0, 2, 0); } else { - x_108 = x_106; + x_107 = x_105; } -lean_ctor_set(x_108, 0, x_107); -lean_ctor_set(x_108, 1, x_105); -return x_108; +lean_ctor_set(x_107, 0, x_106); +lean_ctor_set(x_107, 1, x_104); +return x_107; } else { -lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_109 = lean_ctor_get(x_103, 0); +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; +x_108 = lean_ctor_get(x_102, 0); +lean_inc(x_108); +x_109 = lean_ctor_get(x_102, 1); lean_inc(x_109); -x_110 = lean_ctor_get(x_103, 1); -lean_inc(x_110); -if (lean_is_exclusive(x_103)) { - lean_ctor_release(x_103, 0); - lean_ctor_release(x_103, 1); - x_111 = x_103; +if (lean_is_exclusive(x_102)) { + lean_ctor_release(x_102, 0); + lean_ctor_release(x_102, 1); + x_110 = x_102; } else { - lean_dec_ref(x_103); - x_111 = lean_box(0); + lean_dec_ref(x_102); + x_110 = lean_box(0); } -if (lean_is_scalar(x_111)) { - x_112 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_110)) { + x_111 = lean_alloc_ctor(1, 2, 0); } else { - x_112 = x_111; + x_111 = x_110; } -lean_ctor_set(x_112, 0, x_109); -lean_ctor_set(x_112, 1, x_110); -return x_112; +lean_ctor_set(x_111, 0, x_108); +lean_ctor_set(x_111, 1, x_109); +return x_111; } } } @@ -31490,7 +31493,7 @@ return x_112; } else { -uint8_t x_113; +uint8_t x_112; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -31500,23 +31503,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_2); lean_dec(x_1); -x_113 = !lean_is_exclusive(x_16); -if (x_113 == 0) +x_112 = !lean_is_exclusive(x_15); +if (x_112 == 0) { -return x_16; +return x_15; } else { -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_16, 0); -x_115 = lean_ctor_get(x_16, 1); -lean_inc(x_115); +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_15, 0); +x_114 = lean_ctor_get(x_15, 1); lean_inc(x_114); -lean_dec(x_16); -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_inc(x_113); +lean_dec(x_15); +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; } } } diff --git a/stage0/stdlib/Lean/Elab/MutualDef.c b/stage0/stdlib/Lean/Elab/MutualDef.c index 47ed8fe74053..7d8fd4057068 100644 --- a/stage0/stdlib/Lean/Elab/MutualDef.c +++ b/stage0/stdlib/Lean/Elab/MutualDef.c @@ -4579,56 +4579,57 @@ return x_79; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___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_object* x_18) { _start: { -uint8_t x_19; lean_object* x_20; +uint8_t x_19; uint8_t x_20; lean_object* x_21; lean_dec(x_10); -x_19 = 0; +x_19 = 1; +x_20 = 0; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_20 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_19, x_19, x_12, x_13, x_14, x_15, x_16, x_17, x_18); -if (lean_obj_tag(x_20) == 0) +x_21 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_19, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; uint8_t x_22; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); +lean_object* x_22; uint8_t x_23; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); lean_inc(x_2); -x_22 = l_Lean_Elab_DefView_isInstance(x_2); -if (x_22 == 0) +x_23 = l_Lean_Elab_DefView_isInstance(x_2); +if (x_23 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_box(0); -x_24 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_23, x_12, x_13, x_14, x_15, x_16, x_17, x_21); -return x_24; +lean_object* x_24; lean_object* x_25; +x_24 = lean_box(0); +x_25 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_11, x_24, x_12, x_13, x_14, x_15, x_16, x_17, x_22); +return x_25; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__1; -x_26 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__2; +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__1; +x_27 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_cleanupOfNat___closed__2; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_27 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_11, x_25, x_26, x_19, x_19, x_14, x_15, x_16, x_17, x_21); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_11, x_26, x_27, x_20, x_20, x_14, x_15, x_16, x_17, x_22); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = lean_box(0); -x_31 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_28, x_30, x_12, x_13, x_14, x_15, x_16, x_17, x_29); -return x_31; +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_box(0); +x_32 = l_Array_forInUnsafe_loop___at___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabHeaders___spec__6___lambda__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_29, x_31, x_12, x_13, x_14, x_15, x_16, x_17, x_30); +return x_32; } else { -uint8_t x_32; +uint8_t x_33; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -4644,30 +4645,30 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +x_33 = !lean_is_exclusive(x_28); +if (x_33 == 0) { -return x_27; +return x_28; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 0); +x_35 = lean_ctor_get(x_28, 1); +lean_inc(x_35); lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_dec(x_28); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } } else { -uint8_t x_36; +uint8_t x_37; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -4684,23 +4685,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_20); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_21); +if (x_37 == 0) { -return x_20; +return x_21; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_20, 0); -x_38 = lean_ctor_get(x_20, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_21, 0); +x_39 = lean_ctor_get(x_21, 1); +lean_inc(x_39); lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_20); -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_dec(x_21); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; } } } @@ -29634,64 +29635,65 @@ lean_inc(x_2); x_61 = l___private_Lean_Elab_MutualDef_0__Lean_Elab_Term_elabFunValues(x_2, x_7, x_8, x_9, x_10, x_11, x_12, x_21); if (lean_obj_tag(x_61) == 0) { -lean_object* x_62; lean_object* x_63; uint8_t x_64; lean_object* x_65; +lean_object* x_62; lean_object* x_63; uint8_t x_64; uint8_t 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 = 0; +x_64 = 1; +x_65 = 0; 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_65 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_64, x_64, x_7, x_8, x_9, x_10, x_11, x_12, x_63); -if (lean_obj_tag(x_65) == 0) +x_66 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_64, x_65, x_7, x_8, x_9, x_10, x_11, x_12, x_63); +if (lean_obj_tag(x_66) == 0) { -lean_object* x_66; lean_object* x_67; size_t x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; +lean_object* x_67; lean_object* x_68; size_t x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; lean_dec(x_22); -x_66 = lean_ctor_get(x_65, 1); -lean_inc(x_66); -lean_dec(x_65); -x_67 = lean_array_get_size(x_62); -x_68 = lean_usize_of_nat(x_67); -lean_dec(x_67); -x_69 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(x_68, x_19, x_62, x_7, x_8, x_9, x_10, x_11, x_12, x_66); -x_70 = lean_ctor_get(x_69, 0); -lean_inc(x_70); -x_71 = lean_ctor_get(x_69, 1); +x_67 = lean_ctor_get(x_66, 1); +lean_inc(x_67); +lean_dec(x_66); +x_68 = lean_array_get_size(x_62); +x_69 = lean_usize_of_nat(x_68); +lean_dec(x_68); +x_70 = l_Array_mapMUnsafe_map___at_Lean_Elab_Term_MutualClosure_main___spec__2(x_69, x_19, x_62, x_7, x_8, x_9, x_10, x_11, x_12, x_67); +x_71 = lean_ctor_get(x_70, 0); lean_inc(x_71); -lean_dec(x_69); -x_72 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_19, x_6, x_3, x_4, x_1, x_5, x_70, x_7, x_8, x_9, x_10, x_11, x_12, x_71); +x_72 = lean_ctor_get(x_70, 1); +lean_inc(x_72); +lean_dec(x_70); +x_73 = l_Lean_Elab_Term_elabMutualDef_go___lambda__2(x_2, x_19, x_6, x_3, x_4, x_1, x_5, x_71, x_7, x_8, x_9, x_10, x_11, x_12, x_72); lean_dec(x_5); -return x_72; +return x_73; } else { -lean_object* x_73; lean_object* x_74; +lean_object* x_74; lean_object* x_75; lean_dec(x_62); -x_73 = lean_ctor_get(x_65, 0); -lean_inc(x_73); -x_74 = lean_ctor_get(x_65, 1); +x_74 = lean_ctor_get(x_66, 0); lean_inc(x_74); -lean_dec(x_65); -x_23 = x_73; -x_24 = x_74; +x_75 = lean_ctor_get(x_66, 1); +lean_inc(x_75); +lean_dec(x_66); +x_23 = x_74; +x_24 = x_75; goto block_60; } } else { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_61, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_61, 1); +lean_object* x_76; lean_object* x_77; +x_76 = lean_ctor_get(x_61, 0); lean_inc(x_76); +x_77 = lean_ctor_get(x_61, 1); +lean_inc(x_77); lean_dec(x_61); -x_23 = x_75; -x_24 = x_76; +x_23 = x_76; +x_24 = x_77; goto block_60; } block_60: @@ -29948,7 +29950,7 @@ return x_59; } else { -uint8_t x_77; +uint8_t x_78; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -29961,23 +29963,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_77 = !lean_is_exclusive(x_20); -if (x_77 == 0) +x_78 = !lean_is_exclusive(x_20); +if (x_78 == 0) { return x_20; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_78 = lean_ctor_get(x_20, 0); -x_79 = lean_ctor_get(x_20, 1); +lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_79 = lean_ctor_get(x_20, 0); +x_80 = lean_ctor_get(x_20, 1); +lean_inc(x_80); lean_inc(x_79); -lean_inc(x_78); lean_dec(x_20); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c index a9b9715de7b3..ef4966a9b6ff 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Eqns.c @@ -22,7 +22,6 @@ lean_object* l_Lean_Meta_getMatcherInfo_x3f___at_Lean_Meta_Simp_simpMatch___spec 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*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__11; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo; @@ -31,6 +30,7 @@ lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, l static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__13; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__3___closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__7; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_getEqnsFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_registerEqnsInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*); @@ -44,13 +44,14 @@ LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpM lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113_(lean_object*); uint8_t l_Lean_Expr_isAppOfArity(lean_object*, lean_object*, lean_object*); static lean_object* l_panic___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1___closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__1; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13; lean_object* l_Lean_Expr_sort___override(lean_object*); lean_object* l_Lean_MessageData_ofList(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__3; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Array_toSubarray___rarg(lean_object*, lean_object*, lean_object*); @@ -61,10 +62,8 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__11; lean_object* lean_mk_array(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9; lean_object* l_Lean_MVarId_getType_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__6; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__10; @@ -86,12 +85,13 @@ lean_object* l_Lean_Meta_registerGetUnfoldEqnFn(lean_object*, lean_object*); lean_object* l_Lean_Elab_Eqns_tryContradiction(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__8; uint8_t l_Lean_Expr_isAppOf(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7; lean_object* l_Lean_stringToMessageData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_maxRecDepth; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__8; 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_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__4; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_WF_mkEqns___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appArg_x21(lean_object*); @@ -103,6 +103,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Ela LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_instInhabitedPUnit; uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_WF_registerEqnsInfo___spec__3___boxed(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*); @@ -116,10 +117,7 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__4; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__10; lean_object* lean_st_ref_take(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix(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_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__8; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__4; @@ -135,6 +133,7 @@ lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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_WF_simpMatchWF_x3f___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17; lean_object* l_Lean_Name_getPrefix(lean_object*); lean_object* l_List_mapTR_loop___at_Lean_MessageData_instCoeListExpr___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -144,22 +143,26 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_getUnfoldFor_x3f___lambda__1___boxed(lea LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MVarId_withContext___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Match_MatcherInfo_numAlts(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__6; +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383_(lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__4; lean_object* l_Lean_Meta_mkEqTrans(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15; lean_object* l_Lean_Elab_Eqns_getUnfoldFor_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__9; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___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_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_mkEqns___closed__1; extern lean_object* l_Lean_instInhabitedExpr; @@ -167,7 +170,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Ela lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_WF_registerEqnsInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -176,6 +178,7 @@ uint8_t lean_name_eq(lean_object*, lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__3; lean_object* l_Lean_MapDeclarationExtension_find_x3f___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__1(lean_object*, size_t, size_t); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_mkEqns___closed__2; @@ -186,23 +189,22 @@ static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_ extern lean_object* l_Lean_diagnostics; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___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_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6; LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); lean_object* l_Lean_Elab_Eqns_simpIf_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_instMonadMetaM; lean_object* l_Lean_Elab_Eqns_mkEqnTypes(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__4; lean_object* l_Lean_Meta_splitTarget_x3f(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instBEqOrigin___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; lean_object* l_Lean_Meta_SplitIf_mkDischarge_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); lean_object* l_Lean_Meta_Split_getSimpMatchContext___rarg(lean_object*, lean_object*); @@ -210,6 +212,7 @@ static lean_object* l_panic___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__2___cl static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__3___closed__1; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__12; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2; extern lean_object* l_Lean_Meta_eqnThmSuffixBase; lean_object* l_Array_extract___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns___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*); @@ -217,7 +220,6 @@ lean_object* lean_name_append_index_after(lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__6; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__3; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__8; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12; static lean_object* l_Lean_Elab_WF_mkEqns___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_WF_registerEqnsInfo___spec__2(size_t, size_t, lean_object*); lean_object* l_Lean_Expr_constLevels_x21(lean_object*); @@ -228,18 +230,16 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__3___boxed(lean LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__12; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11; lean_object* l_Lean_instantiateMVars___at___private_Lean_Meta_Basic_0__Lean_Meta_isClassApp_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__2; static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__14; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__8; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__9; @@ -251,6 +251,7 @@ static lean_object* l_panic___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__2___cl static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__7; static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__5; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__1; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__5; @@ -259,7 +260,6 @@ static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__15; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__4; extern lean_object* l_Lean_Expr_instHashable; static lean_object* l_Lean_Elab_WF_simpMatchWF_x3f___lambda__4___closed__1; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Expr_getAppFn(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__1; @@ -268,22 +268,23 @@ lean_object* l_Lean_mkMapDeclarationExtension___rarg(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_WF_registerEqnsInfo___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___closed__1; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__6; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18; lean_object* l_Lean_Meta_ensureEqnReservedNamesAvailable(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_rwFixEq___lambda__1___closed__4; static lean_object* l_Lean_Elab_WF_registerEqnsInfo___closed__2; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___closed__2; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_deltaLHSUntilFix___lambda__1___closed__9; -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2; lean_object* l_Lean_Elab_Eqns_whnfReducibleLHS_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_Meta_lambdaTelescope___at_Lean_PrettyPrinter_Delaborator_returnsPi___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof___closed__1; +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5; +uint8_t l_Lean_Meta_TransparencyMode_lt(uint8_t, uint8_t); size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_instHashableOrigin___boxed(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5; static lean_object* l_List_forIn_loop___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__3___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; @@ -292,6 +293,7 @@ lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Lean_Meta_casesOnStuckLHS_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Eqns_removeUnusedEqnHypotheses(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Meta_RecursorInfo_0__Lean_Meta_getMajorPosDepElim___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4; lean_object* l_instInhabitedOfMonad___rarg(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_simpTargetStar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -311,11 +313,11 @@ lean_object* lean_infer_type(lean_object*, lean_object*, lean_object*, lean_obje uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__3; uint8_t lean_usize_dec_lt(size_t, size_t); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3; static lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; lean_object* l_Lean_MVarId_intros(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_panic___at_Lean_Elab_WF_simpMatchWF_x3f_pre___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_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_WF_mkEqns(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_applySimpResultToTarget(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_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__5; @@ -325,12 +327,11 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Ela lean_object* l_Lean_Expr_mvarId_x21(lean_object*); static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__2; lean_object* l_Lean_InductiveVal_numCtors(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14; static lean_object* l_Lean_Elab_WF_instInhabitedEqnInfo___closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); -static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15; LEAN_EXPORT lean_object* l_Lean_Meta_matchMatcherApp_x3f___at_Lean_Elab_WF_simpMatchWF_x3f_pre___spec__1___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_panic___at_Lean_Meta_subst_substEq___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_addDecl(lean_object*, lean_object*, lean_object*, lean_object*); @@ -3812,378 +3813,681 @@ return x_2; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___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; +uint8_t x_9; lean_object* x_10; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; lean_dec(x_3); +x_135 = lean_ctor_get(x_4, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_4, 1); +lean_inc(x_136); +x_137 = lean_ctor_get(x_4, 2); +lean_inc(x_137); +x_138 = lean_ctor_get(x_4, 3); +lean_inc(x_138); +x_139 = lean_ctor_get(x_4, 4); +lean_inc(x_139); +x_140 = lean_ctor_get(x_4, 5); +lean_inc(x_140); +x_141 = !lean_is_exclusive(x_135); +if (x_141 == 0) +{ +uint8_t x_142; uint8_t x_143; uint8_t x_144; uint8_t x_145; uint8_t x_146; +x_142 = lean_ctor_get_uint8(x_4, sizeof(void*)*6); +x_143 = lean_ctor_get_uint8(x_4, sizeof(void*)*6 + 1); +x_144 = lean_ctor_get_uint8(x_135, 9); +x_145 = 0; +x_146 = l_Lean_Meta_TransparencyMode_lt(x_144, x_145); +if (x_146 == 0) +{ +lean_object* x_147; lean_object* x_148; +x_147 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_147, 0, x_135); +lean_ctor_set(x_147, 1, x_136); +lean_ctor_set(x_147, 2, x_137); +lean_ctor_set(x_147, 3, x_138); +lean_ctor_set(x_147, 4, x_139); +lean_ctor_set(x_147, 5, x_140); +lean_ctor_set_uint8(x_147, sizeof(void*)*6, x_142); +lean_ctor_set_uint8(x_147, sizeof(void*)*6 + 1, x_143); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -lean_inc(x_4); lean_inc(x_1); -x_9 = l_Lean_Elab_Eqns_tryURefl(x_1, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_9) == 0) +x_148 = l_Lean_Elab_Eqns_tryURefl(x_1, x_147, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_148) == 0) +{ +lean_object* x_149; lean_object* x_150; uint8_t x_151; +x_149 = lean_ctor_get(x_148, 0); +lean_inc(x_149); +x_150 = lean_ctor_get(x_148, 1); +lean_inc(x_150); +lean_dec(x_148); +x_151 = lean_unbox(x_149); +lean_dec(x_149); +x_9 = x_151; +x_10 = x_150; +goto block_134; +} +else { -lean_object* x_10; uint8_t x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -x_11 = lean_unbox(x_10); -lean_dec(x_10); -if (x_11 == 0) +uint8_t x_152; +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_152 = !lean_is_exclusive(x_148); +if (x_152 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_9, 1); -lean_inc(x_12); -lean_dec(x_9); +return x_148; +} +else +{ +lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_153 = lean_ctor_get(x_148, 0); +x_154 = lean_ctor_get(x_148, 1); +lean_inc(x_154); +lean_inc(x_153); +lean_dec(x_148); +x_155 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_155, 0, x_153); +lean_ctor_set(x_155, 1, x_154); +return x_155; +} +} +} +else +{ +lean_object* x_156; lean_object* x_157; +lean_ctor_set_uint8(x_135, 9, x_145); +x_156 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_156, 0, x_135); +lean_ctor_set(x_156, 1, x_136); +lean_ctor_set(x_156, 2, x_137); +lean_ctor_set(x_156, 3, x_138); +lean_ctor_set(x_156, 4, x_139); +lean_ctor_set(x_156, 5, x_140); +lean_ctor_set_uint8(x_156, sizeof(void*)*6, x_142); +lean_ctor_set_uint8(x_156, sizeof(void*)*6 + 1, x_143); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_157 = l_Lean_Elab_Eqns_tryURefl(x_1, x_156, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_157) == 0) +{ +lean_object* x_158; lean_object* x_159; uint8_t x_160; +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_unbox(x_158); +lean_dec(x_158); +x_9 = x_160; +x_10 = x_159; +goto block_134; +} +else +{ +uint8_t x_161; +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_161 = !lean_is_exclusive(x_157); +if (x_161 == 0) +{ +return x_157; +} +else +{ +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_157, 0); +x_163 = lean_ctor_get(x_157, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_157); +x_164 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +return x_164; +} +} +} +} +else +{ +uint8_t x_165; uint8_t x_166; uint8_t x_167; uint8_t x_168; uint8_t x_169; uint8_t x_170; uint8_t x_171; uint8_t x_172; uint8_t x_173; uint8_t x_174; uint8_t x_175; uint8_t x_176; uint8_t x_177; uint8_t x_178; uint8_t x_179; uint8_t x_180; uint8_t x_181; +x_165 = lean_ctor_get_uint8(x_4, sizeof(void*)*6); +x_166 = lean_ctor_get_uint8(x_4, sizeof(void*)*6 + 1); +x_167 = lean_ctor_get_uint8(x_135, 0); +x_168 = lean_ctor_get_uint8(x_135, 1); +x_169 = lean_ctor_get_uint8(x_135, 2); +x_170 = lean_ctor_get_uint8(x_135, 3); +x_171 = lean_ctor_get_uint8(x_135, 4); +x_172 = lean_ctor_get_uint8(x_135, 5); +x_173 = lean_ctor_get_uint8(x_135, 6); +x_174 = lean_ctor_get_uint8(x_135, 7); +x_175 = lean_ctor_get_uint8(x_135, 8); +x_176 = lean_ctor_get_uint8(x_135, 9); +x_177 = lean_ctor_get_uint8(x_135, 10); +x_178 = lean_ctor_get_uint8(x_135, 11); +x_179 = lean_ctor_get_uint8(x_135, 12); +lean_dec(x_135); +x_180 = 0; +x_181 = l_Lean_Meta_TransparencyMode_lt(x_176, x_180); +if (x_181 == 0) +{ +lean_object* x_182; lean_object* x_183; lean_object* x_184; +x_182 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_182, 0, x_167); +lean_ctor_set_uint8(x_182, 1, x_168); +lean_ctor_set_uint8(x_182, 2, x_169); +lean_ctor_set_uint8(x_182, 3, x_170); +lean_ctor_set_uint8(x_182, 4, x_171); +lean_ctor_set_uint8(x_182, 5, x_172); +lean_ctor_set_uint8(x_182, 6, x_173); +lean_ctor_set_uint8(x_182, 7, x_174); +lean_ctor_set_uint8(x_182, 8, x_175); +lean_ctor_set_uint8(x_182, 9, x_176); +lean_ctor_set_uint8(x_182, 10, x_177); +lean_ctor_set_uint8(x_182, 11, x_178); +lean_ctor_set_uint8(x_182, 12, x_179); +x_183 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_183, 0, x_182); +lean_ctor_set(x_183, 1, x_136); +lean_ctor_set(x_183, 2, x_137); +lean_ctor_set(x_183, 3, x_138); +lean_ctor_set(x_183, 4, x_139); +lean_ctor_set(x_183, 5, x_140); +lean_ctor_set_uint8(x_183, sizeof(void*)*6, x_165); +lean_ctor_set_uint8(x_183, sizeof(void*)*6 + 1, x_166); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_184 = l_Lean_Elab_Eqns_tryURefl(x_1, x_183, x_5, x_6, x_7, x_8); +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 = lean_unbox(x_185); +lean_dec(x_185); +x_9 = x_187; +x_10 = x_186; +goto block_134; +} +else +{ +lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +lean_dec(x_1); +x_188 = lean_ctor_get(x_184, 0); +lean_inc(x_188); +x_189 = lean_ctor_get(x_184, 1); +lean_inc(x_189); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_190 = x_184; +} else { + lean_dec_ref(x_184); + 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; +} +} +else +{ +lean_object* x_192; lean_object* x_193; lean_object* x_194; +x_192 = lean_alloc_ctor(0, 0, 13); +lean_ctor_set_uint8(x_192, 0, x_167); +lean_ctor_set_uint8(x_192, 1, x_168); +lean_ctor_set_uint8(x_192, 2, x_169); +lean_ctor_set_uint8(x_192, 3, x_170); +lean_ctor_set_uint8(x_192, 4, x_171); +lean_ctor_set_uint8(x_192, 5, x_172); +lean_ctor_set_uint8(x_192, 6, x_173); +lean_ctor_set_uint8(x_192, 7, x_174); +lean_ctor_set_uint8(x_192, 8, x_175); +lean_ctor_set_uint8(x_192, 9, x_180); +lean_ctor_set_uint8(x_192, 10, x_177); +lean_ctor_set_uint8(x_192, 11, x_178); +lean_ctor_set_uint8(x_192, 12, x_179); +x_193 = lean_alloc_ctor(0, 6, 2); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_136); +lean_ctor_set(x_193, 2, x_137); +lean_ctor_set(x_193, 3, x_138); +lean_ctor_set(x_193, 4, x_139); +lean_ctor_set(x_193, 5, x_140); +lean_ctor_set_uint8(x_193, sizeof(void*)*6, x_165); +lean_ctor_set_uint8(x_193, sizeof(void*)*6 + 1, x_166); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_1); +x_194 = l_Lean_Elab_Eqns_tryURefl(x_1, x_193, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_194) == 0) +{ +lean_object* x_195; lean_object* x_196; uint8_t x_197; +x_195 = lean_ctor_get(x_194, 0); +lean_inc(x_195); +x_196 = lean_ctor_get(x_194, 1); +lean_inc(x_196); +lean_dec(x_194); +x_197 = lean_unbox(x_195); +lean_dec(x_195); +x_9 = x_197; +x_10 = x_196; +goto block_134; +} +else +{ +lean_object* x_198; lean_object* x_199; lean_object* x_200; lean_object* x_201; +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_198 = lean_ctor_get(x_194, 0); +lean_inc(x_198); +x_199 = lean_ctor_get(x_194, 1); +lean_inc(x_199); +if (lean_is_exclusive(x_194)) { + lean_ctor_release(x_194, 0); + lean_ctor_release(x_194, 1); + x_200 = x_194; +} else { + lean_dec_ref(x_194); + x_200 = lean_box(0); +} +if (lean_is_scalar(x_200)) { + x_201 = lean_alloc_ctor(1, 2, 0); +} else { + x_201 = x_200; +} +lean_ctor_set(x_201, 0, x_198); +lean_ctor_set(x_201, 1, x_199); +return x_201; +} +} +} +block_134: +{ +if (x_9 == 0) +{ +lean_object* x_11; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_13 = l_Lean_Elab_Eqns_tryContradiction(x_1, x_4, x_5, x_6, x_7, x_12); -if (lean_obj_tag(x_13) == 0) +x_11 = l_Lean_Elab_Eqns_tryContradiction(x_1, x_4, x_5, x_6, x_7, x_10); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_14; uint8_t x_15; -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_unbox(x_14); -lean_dec(x_14); -if (x_15 == 0) +lean_object* x_12; uint8_t x_13; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); +x_13 = lean_unbox(x_12); +lean_dec(x_12); +if (x_13 == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -lean_dec(x_13); +lean_object* x_14; lean_object* x_15; +x_14 = lean_ctor_get(x_11, 1); +lean_inc(x_14); +lean_dec(x_11); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_17 = l_Lean_Elab_WF_simpMatchWF_x3f(x_1, x_4, x_5, x_6, x_7, x_16); -if (lean_obj_tag(x_17) == 0) +x_15 = l_Lean_Elab_WF_simpMatchWF_x3f(x_1, x_4, x_5, x_6, x_7, x_14); +if (lean_obj_tag(x_15) == 0) { -lean_object* x_18; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -if (lean_obj_tag(x_18) == 0) +lean_object* x_16; +x_16 = lean_ctor_get(x_15, 0); +lean_inc(x_16); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); +lean_object* x_17; lean_object* x_18; +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_1); -x_20 = l_Lean_Elab_Eqns_simpIf_x3f(x_1, x_4, x_5, x_6, x_7, x_19); -if (lean_obj_tag(x_20) == 0) +x_18 = l_Lean_Elab_Eqns_simpIf_x3f(x_1, x_4, x_5, x_6, x_7, x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_21; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -if (lean_obj_tag(x_21) == 0) +lean_object* x_19; +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_23 = l_Lean_Elab_Eqns_whnfReducibleLHS_x3f(x_1, x_4, x_5, x_6, x_7, x_22); -if (lean_obj_tag(x_23) == 0) +x_21 = l_Lean_Elab_Eqns_whnfReducibleLHS_x3f(x_1, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_24; -x_24 = lean_ctor_get(x_23, 0); -lean_inc(x_24); -if (lean_obj_tag(x_24) == 0) +lean_object* x_22; +x_22 = lean_ctor_get(x_21, 0); +lean_inc(x_22); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_25 = lean_ctor_get(x_23, 1); -lean_inc(x_25); -lean_dec(x_23); -x_26 = lean_box(0); -x_27 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; -x_28 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; -x_29 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_23 = lean_ctor_get(x_21, 1); +lean_inc(x_23); +lean_dec(x_21); +x_24 = lean_box(0); +x_25 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__5; +x_26 = l_Lean_Elab_WF_instInhabitedEqnInfo___closed__3; +x_27 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__7; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_30 = l_Lean_Meta_simpTargetStar(x_1, x_27, x_28, x_26, x_29, x_4, x_5, x_6, x_7, x_25); -if (lean_obj_tag(x_30) == 0) +x_28 = l_Lean_Meta_simpTargetStar(x_1, x_25, x_26, x_24, x_27, x_4, x_5, x_6, x_7, x_23); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_31, 0); -lean_inc(x_32); -lean_dec(x_31); -switch (lean_obj_tag(x_32)) { +lean_object* x_29; lean_object* x_30; +x_29 = lean_ctor_get(x_28, 0); +lean_inc(x_29); +x_30 = lean_ctor_get(x_29, 0); +lean_inc(x_30); +lean_dec(x_29); +switch (lean_obj_tag(x_30)) { case 0: { -uint8_t x_33; +uint8_t x_31; 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_33 = !lean_is_exclusive(x_30); -if (x_33 == 0) +x_31 = !lean_is_exclusive(x_28); +if (x_31 == 0) { -lean_object* x_34; lean_object* x_35; -x_34 = lean_ctor_get(x_30, 0); -lean_dec(x_34); -x_35 = lean_box(0); -lean_ctor_set(x_30, 0, x_35); -return x_30; +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_28, 0); +lean_dec(x_32); +x_33 = lean_box(0); +lean_ctor_set(x_28, 0, x_33); +return x_28; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_dec(x_30); -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; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_28, 1); +lean_inc(x_34); +lean_dec(x_28); +x_35 = lean_box(0); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_35); +lean_ctor_set(x_36, 1, x_34); +return x_36; } } case 1: { -lean_object* x_39; lean_object* x_40; -x_39 = lean_ctor_get(x_30, 1); -lean_inc(x_39); -lean_dec(x_30); +lean_object* x_37; lean_object* x_38; +x_37 = lean_ctor_get(x_28, 1); +lean_inc(x_37); +lean_dec(x_28); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_40 = l_Lean_Meta_casesOnStuckLHS_x3f(x_1, x_4, x_5, x_6, x_7, x_39); -if (lean_obj_tag(x_40) == 0) +x_38 = l_Lean_Meta_casesOnStuckLHS_x3f(x_1, x_4, x_5, x_6, x_7, x_37); +if (lean_obj_tag(x_38) == 0) { -lean_object* x_41; -x_41 = lean_ctor_get(x_40, 0); -lean_inc(x_41); -if (lean_obj_tag(x_41) == 0) +lean_object* x_39; +x_39 = lean_ctor_get(x_38, 0); +lean_inc(x_39); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_42; uint8_t x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_40, 1); -lean_inc(x_42); -lean_dec(x_40); -x_43 = 1; +lean_object* x_40; uint8_t x_41; lean_object* x_42; +x_40 = lean_ctor_get(x_38, 1); +lean_inc(x_40); +lean_dec(x_38); +x_41 = 1; lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_1); -x_44 = l_Lean_Meta_splitTarget_x3f(x_1, x_43, x_4, x_5, x_6, x_7, x_42); -if (lean_obj_tag(x_44) == 0) +x_42 = l_Lean_Meta_splitTarget_x3f(x_1, x_41, x_4, x_5, x_6, x_7, x_40); +if (lean_obj_tag(x_42) == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_44, 0); -lean_inc(x_45); -if (lean_obj_tag(x_45) == 0) +lean_object* x_43; +x_43 = lean_ctor_get(x_42, 0); +lean_inc(x_43); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_46 = lean_ctor_get(x_44, 1); -lean_inc(x_46); -lean_dec(x_44); -x_47 = l_Lean_MessageData_ofName(x_2); -x_48 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___closed__4; +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; 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_42, 1); +lean_inc(x_44); +lean_dec(x_42); +x_45 = l_Lean_MessageData_ofName(x_2); +x_46 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__3___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___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9; x_49 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_49, 0, x_48); -lean_ctor_set(x_49, 1, x_47); -x_50 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___lambda__1___closed__9; +lean_ctor_set(x_49, 0, x_47); +lean_ctor_set(x_49, 1, x_48); +x_50 = lean_alloc_ctor(2, 1, 0); +lean_ctor_set(x_50, 0, x_1); x_51 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_51, 0, x_49); lean_ctor_set(x_51, 1, x_50); -x_52 = lean_alloc_ctor(2, 1, 0); -lean_ctor_set(x_52, 0, x_1); +x_52 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__9; x_53 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_53, 0, x_51); lean_ctor_set(x_53, 1, x_52); -x_54 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__9; -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_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_55, x_4, x_5, x_6, x_7, x_46); +x_54 = l_Lean_throwError___at_Lean_Meta_setInlineAttribute___spec__1(x_53, x_4, x_5, x_6, x_7, x_44); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_56; +return x_54; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; +lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_dec(x_1); -x_57 = lean_ctor_get(x_44, 1); -lean_inc(x_57); -lean_dec(x_44); -x_58 = lean_ctor_get(x_45, 0); -lean_inc(x_58); -lean_dec(x_45); -x_59 = l_List_forM___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__1(x_2, x_58, x_4, x_5, x_6, x_7, x_57); -return x_59; +x_55 = lean_ctor_get(x_42, 1); +lean_inc(x_55); +lean_dec(x_42); +x_56 = lean_ctor_get(x_43, 0); +lean_inc(x_56); +lean_dec(x_43); +x_57 = l_List_forM___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__1(x_2, x_56, x_4, x_5, x_6, x_7, x_55); +return x_57; } } else { -uint8_t x_60; +uint8_t x_58; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_60 = !lean_is_exclusive(x_44); -if (x_60 == 0) +x_58 = !lean_is_exclusive(x_42); +if (x_58 == 0) { -return x_44; +return x_42; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_44, 0); -x_62 = lean_ctor_get(x_44, 1); -lean_inc(x_62); -lean_inc(x_61); -lean_dec(x_44); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_61); -lean_ctor_set(x_63, 1, x_62); -return x_63; +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_42, 0); +x_60 = lean_ctor_get(x_42, 1); +lean_inc(x_60); +lean_inc(x_59); +lean_dec(x_42); +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 { -uint8_t x_64; +uint8_t x_62; lean_dec(x_1); -x_64 = !lean_is_exclusive(x_40); -if (x_64 == 0) +x_62 = !lean_is_exclusive(x_38); +if (x_62 == 0) { -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_40, 1); -x_66 = lean_ctor_get(x_40, 0); -lean_dec(x_66); -x_67 = lean_ctor_get(x_41, 0); -lean_inc(x_67); -lean_dec(x_41); -x_68 = lean_array_get_size(x_67); -x_69 = lean_unsigned_to_nat(0u); -x_70 = lean_nat_dec_lt(x_69, x_68); -if (x_70 == 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_ctor_get(x_38, 1); +x_64 = lean_ctor_get(x_38, 0); +lean_dec(x_64); +x_65 = lean_ctor_get(x_39, 0); +lean_inc(x_65); +lean_dec(x_39); +x_66 = lean_array_get_size(x_65); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +if (x_68 == 0) { -lean_object* x_71; -lean_dec(x_68); -lean_dec(x_67); +lean_object* x_69; +lean_dec(x_66); +lean_dec(x_65); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_71 = lean_box(0); -lean_ctor_set(x_40, 0, x_71); -return x_40; +x_69 = lean_box(0); +lean_ctor_set(x_38, 0, x_69); +return x_38; } else { -uint8_t x_72; -x_72 = lean_nat_dec_le(x_68, x_68); -if (x_72 == 0) +uint8_t x_70; +x_70 = lean_nat_dec_le(x_66, x_66); +if (x_70 == 0) { -lean_object* x_73; -lean_dec(x_68); -lean_dec(x_67); +lean_object* x_71; +lean_dec(x_66); +lean_dec(x_65); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_73 = lean_box(0); -lean_ctor_set(x_40, 0, x_73); -return x_40; +x_71 = lean_box(0); +lean_ctor_set(x_38, 0, x_71); +return x_38; } else { -size_t x_74; size_t x_75; lean_object* x_76; lean_object* x_77; -lean_free_object(x_40); -x_74 = 0; -x_75 = lean_usize_of_nat(x_68); -lean_dec(x_68); -x_76 = lean_box(0); -x_77 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(x_2, x_67, x_74, x_75, x_76, x_4, x_5, x_6, x_7, x_65); -lean_dec(x_67); -return x_77; +size_t x_72; size_t x_73; lean_object* x_74; lean_object* x_75; +lean_free_object(x_38); +x_72 = 0; +x_73 = lean_usize_of_nat(x_66); +lean_dec(x_66); +x_74 = lean_box(0); +x_75 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(x_2, x_65, x_72, x_73, x_74, x_4, x_5, x_6, x_7, x_63); +lean_dec(x_65); +return x_75; } } } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; uint8_t x_82; -x_78 = lean_ctor_get(x_40, 1); -lean_inc(x_78); -lean_dec(x_40); -x_79 = lean_ctor_get(x_41, 0); -lean_inc(x_79); -lean_dec(x_41); -x_80 = lean_array_get_size(x_79); -x_81 = lean_unsigned_to_nat(0u); -x_82 = lean_nat_dec_lt(x_81, x_80); -if (x_82 == 0) +lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; uint8_t x_80; +x_76 = lean_ctor_get(x_38, 1); +lean_inc(x_76); +lean_dec(x_38); +x_77 = lean_ctor_get(x_39, 0); +lean_inc(x_77); +lean_dec(x_39); +x_78 = lean_array_get_size(x_77); +x_79 = lean_unsigned_to_nat(0u); +x_80 = lean_nat_dec_lt(x_79, x_78); +if (x_80 == 0) { -lean_object* x_83; lean_object* x_84; -lean_dec(x_80); -lean_dec(x_79); +lean_object* x_81; lean_object* x_82; +lean_dec(x_78); +lean_dec(x_77); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_83 = lean_box(0); -x_84 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_84, 0, x_83); -lean_ctor_set(x_84, 1, x_78); -return x_84; +x_81 = lean_box(0); +x_82 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_82, 0, x_81); +lean_ctor_set(x_82, 1, x_76); +return x_82; } else { -uint8_t x_85; -x_85 = lean_nat_dec_le(x_80, x_80); -if (x_85 == 0) +uint8_t x_83; +x_83 = lean_nat_dec_le(x_78, x_78); +if (x_83 == 0) { -lean_object* x_86; lean_object* x_87; -lean_dec(x_80); -lean_dec(x_79); +lean_object* x_84; lean_object* x_85; +lean_dec(x_78); +lean_dec(x_77); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); -x_86 = lean_box(0); -x_87 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_87, 0, x_86); -lean_ctor_set(x_87, 1, x_78); -return x_87; +x_84 = lean_box(0); +x_85 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_85, 0, x_84); +lean_ctor_set(x_85, 1, x_76); +return x_85; } else { -size_t x_88; size_t x_89; lean_object* x_90; lean_object* x_91; -x_88 = 0; -x_89 = lean_usize_of_nat(x_80); -lean_dec(x_80); -x_90 = lean_box(0); -x_91 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(x_2, x_79, x_88, x_89, x_90, x_4, x_5, x_6, x_7, x_78); -lean_dec(x_79); -return x_91; +size_t x_86; size_t x_87; lean_object* x_88; lean_object* x_89; +x_86 = 0; +x_87 = lean_usize_of_nat(x_78); +lean_dec(x_78); +x_88 = lean_box(0); +x_89 = l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go___spec__2(x_2, x_77, x_86, x_87, x_88, x_4, x_5, x_6, x_7, x_76); +lean_dec(x_77); +return x_89; } } } @@ -4191,327 +4495,282 @@ return x_91; } else { -uint8_t x_92; +uint8_t x_90; 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_92 = !lean_is_exclusive(x_40); -if (x_92 == 0) +x_90 = !lean_is_exclusive(x_38); +if (x_90 == 0) { -return x_40; +return x_38; } else { -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_40, 0); -x_94 = lean_ctor_get(x_40, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_40); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -return x_95; +lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_91 = lean_ctor_get(x_38, 0); +x_92 = lean_ctor_get(x_38, 1); +lean_inc(x_92); +lean_inc(x_91); +lean_dec(x_38); +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; } } } default: { -lean_object* x_96; lean_object* x_97; lean_object* x_98; +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_dec(x_1); -x_96 = lean_ctor_get(x_30, 1); -lean_inc(x_96); +x_94 = lean_ctor_get(x_28, 1); +lean_inc(x_94); +lean_dec(x_28); +x_95 = lean_ctor_get(x_30, 0); +lean_inc(x_95); lean_dec(x_30); -x_97 = lean_ctor_get(x_32, 0); -lean_inc(x_97); -lean_dec(x_32); -x_98 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_97, x_4, x_5, x_6, x_7, x_96); -return x_98; +x_96 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_95, x_4, x_5, x_6, x_7, x_94); +return x_96; } } } else { -uint8_t x_99; +uint8_t x_97; 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_99 = !lean_is_exclusive(x_30); -if (x_99 == 0) +x_97 = !lean_is_exclusive(x_28); +if (x_97 == 0) { -return x_30; +return x_28; } else { -lean_object* x_100; lean_object* x_101; lean_object* x_102; -x_100 = lean_ctor_get(x_30, 0); -x_101 = lean_ctor_get(x_30, 1); -lean_inc(x_101); -lean_inc(x_100); -lean_dec(x_30); -x_102 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_102, 0, x_100); -lean_ctor_set(x_102, 1, x_101); -return x_102; +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_28, 0); +x_99 = lean_ctor_get(x_28, 1); +lean_inc(x_99); +lean_inc(x_98); +lean_dec(x_28); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_dec(x_1); -x_103 = lean_ctor_get(x_23, 1); -lean_inc(x_103); -lean_dec(x_23); -x_104 = lean_ctor_get(x_24, 0); -lean_inc(x_104); -lean_dec(x_24); -x_105 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_104, x_4, x_5, x_6, x_7, x_103); -return x_105; +x_101 = lean_ctor_get(x_21, 1); +lean_inc(x_101); +lean_dec(x_21); +x_102 = lean_ctor_get(x_22, 0); +lean_inc(x_102); +lean_dec(x_22); +x_103 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_102, x_4, x_5, x_6, x_7, x_101); +return x_103; } } else { -uint8_t x_106; +uint8_t x_104; 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_106 = !lean_is_exclusive(x_23); -if (x_106 == 0) +x_104 = !lean_is_exclusive(x_21); +if (x_104 == 0) { -return x_23; +return x_21; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_23, 0); -x_108 = lean_ctor_get(x_23, 1); -lean_inc(x_108); -lean_inc(x_107); -lean_dec(x_23); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +lean_object* x_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_21, 0); +x_106 = lean_ctor_get(x_21, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_21); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } } } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; +lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_dec(x_1); -x_110 = lean_ctor_get(x_20, 1); -lean_inc(x_110); -lean_dec(x_20); -x_111 = lean_ctor_get(x_21, 0); -lean_inc(x_111); -lean_dec(x_21); -x_112 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_111, x_4, x_5, x_6, x_7, x_110); -return x_112; +x_108 = lean_ctor_get(x_18, 1); +lean_inc(x_108); +lean_dec(x_18); +x_109 = lean_ctor_get(x_19, 0); +lean_inc(x_109); +lean_dec(x_19); +x_110 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_109, x_4, x_5, x_6, x_7, x_108); +return x_110; } } else { -uint8_t x_113; +uint8_t x_111; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_113 = !lean_is_exclusive(x_20); -if (x_113 == 0) +x_111 = !lean_is_exclusive(x_18); +if (x_111 == 0) { -return x_20; -} -else -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; -x_114 = lean_ctor_get(x_20, 0); -x_115 = lean_ctor_get(x_20, 1); -lean_inc(x_115); -lean_inc(x_114); -lean_dec(x_20); -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; -} -} +return x_18; } else { -lean_object* x_117; lean_object* x_118; lean_object* x_119; -lean_dec(x_1); -x_117 = lean_ctor_get(x_17, 1); -lean_inc(x_117); -lean_dec(x_17); -x_118 = lean_ctor_get(x_18, 0); -lean_inc(x_118); +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_18, 0); +x_113 = lean_ctor_get(x_18, 1); +lean_inc(x_113); +lean_inc(x_112); lean_dec(x_18); -x_119 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_118, x_4, x_5, x_6, x_7, x_117); -return x_119; +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } -else -{ -uint8_t x_120; -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_120 = !lean_is_exclusive(x_17); -if (x_120 == 0) -{ -return x_17; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; -x_121 = lean_ctor_get(x_17, 0); -x_122 = lean_ctor_get(x_17, 1); -lean_inc(x_122); -lean_inc(x_121); -lean_dec(x_17); -x_123 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_123, 0, x_121); -lean_ctor_set(x_123, 1, x_122); -return x_123; -} +lean_object* x_115; lean_object* x_116; lean_object* x_117; +lean_dec(x_1); +x_115 = lean_ctor_get(x_15, 1); +lean_inc(x_115); +lean_dec(x_15); +x_116 = lean_ctor_get(x_16, 0); +lean_inc(x_116); +lean_dec(x_16); +x_117 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_mkProof_go(x_2, x_116, x_4, x_5, x_6, x_7, x_115); +return x_117; } } else { -uint8_t x_124; +uint8_t x_118; 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_124 = !lean_is_exclusive(x_13); -if (x_124 == 0) +x_118 = !lean_is_exclusive(x_15); +if (x_118 == 0) { -lean_object* x_125; lean_object* x_126; -x_125 = lean_ctor_get(x_13, 0); -lean_dec(x_125); -x_126 = lean_box(0); -lean_ctor_set(x_13, 0, x_126); -return x_13; +return x_15; } else { -lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_127 = lean_ctor_get(x_13, 1); -lean_inc(x_127); -lean_dec(x_13); -x_128 = lean_box(0); -x_129 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_127); -return x_129; +lean_object* x_119; lean_object* x_120; lean_object* x_121; +x_119 = lean_ctor_get(x_15, 0); +x_120 = lean_ctor_get(x_15, 1); +lean_inc(x_120); +lean_inc(x_119); +lean_dec(x_15); +x_121 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +return x_121; } } } else { -uint8_t x_130; +uint8_t x_122; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_130 = !lean_is_exclusive(x_13); -if (x_130 == 0) -{ -return x_13; +x_122 = !lean_is_exclusive(x_11); +if (x_122 == 0) +{ +lean_object* x_123; lean_object* x_124; +x_123 = lean_ctor_get(x_11, 0); +lean_dec(x_123); +x_124 = lean_box(0); +lean_ctor_set(x_11, 0, x_124); +return x_11; } else { -lean_object* x_131; lean_object* x_132; lean_object* x_133; -x_131 = lean_ctor_get(x_13, 0); -x_132 = lean_ctor_get(x_13, 1); -lean_inc(x_132); -lean_inc(x_131); -lean_dec(x_13); -x_133 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -return x_133; +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_11, 1); +lean_inc(x_125); +lean_dec(x_11); +x_126 = lean_box(0); +x_127 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_127, 0, x_126); +lean_ctor_set(x_127, 1, x_125); +return x_127; } } } else { -uint8_t x_134; +uint8_t x_128; 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_134 = !lean_is_exclusive(x_9); -if (x_134 == 0) +x_128 = !lean_is_exclusive(x_11); +if (x_128 == 0) { -lean_object* x_135; lean_object* x_136; -x_135 = lean_ctor_get(x_9, 0); -lean_dec(x_135); -x_136 = lean_box(0); -lean_ctor_set(x_9, 0, x_136); -return x_9; +return x_11; } else { -lean_object* x_137; lean_object* x_138; lean_object* x_139; -x_137 = lean_ctor_get(x_9, 1); -lean_inc(x_137); -lean_dec(x_9); -x_138 = lean_box(0); -x_139 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_139, 0, x_138); -lean_ctor_set(x_139, 1, x_137); -return x_139; +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_11, 0); +x_130 = lean_ctor_get(x_11, 1); +lean_inc(x_130); +lean_inc(x_129); +lean_dec(x_11); +x_131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +return x_131; } } } else { -uint8_t x_140; +lean_object* x_132; lean_object* x_133; 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_140 = !lean_is_exclusive(x_9); -if (x_140 == 0) -{ -return x_9; -} -else -{ -lean_object* x_141; lean_object* x_142; lean_object* x_143; -x_141 = lean_ctor_get(x_9, 0); -x_142 = lean_ctor_get(x_9, 1); -lean_inc(x_142); -lean_inc(x_141); -lean_dec(x_9); -x_143 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_143, 0, x_141); -lean_ctor_set(x_143, 1, x_142); -return x_143; +x_132 = lean_box(0); +x_133 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_133, 0, x_132); +lean_ctor_set(x_133, 1, x_10); +return x_133; } } } @@ -6108,7 +6367,7 @@ lean_dec(x_7); return x_13; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1() { _start: { lean_object* x_1; @@ -6116,7 +6375,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2() { _start: { lean_object* x_1; @@ -6124,7 +6383,7 @@ x_1 = lean_mk_string_from_bytes("WF", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3() { _start: { lean_object* x_1; @@ -6132,23 +6391,23 @@ x_1 = lean_mk_string_from_bytes("eqnInfoExt", 10); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____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_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; -x_3 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; -x_4 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3; +x_3 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2; +x_4 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__4; x_3 = l_Lean_mkMapDeclarationExtension___rarg(x_2, x_1); return x_3; } @@ -7384,7 +7643,7 @@ lean_dec(x_3); return x_4; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1() { _start: { lean_object* x_1; @@ -7392,7 +7651,7 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_getEqnsFor_x3f), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2() { _start: { lean_object* x_1; @@ -7400,37 +7659,37 @@ x_1 = lean_alloc_closure((void*)(l_Lean_Elab_WF_getUnfoldFor_x3f), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6() { _start: { lean_object* x_1; @@ -7438,17 +7697,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8() { _start: { lean_object* x_1; @@ -7456,37 +7715,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10; x_2 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12() { _start: { lean_object* x_1; @@ -7494,27 +7753,27 @@ x_1 = lean_mk_string_from_bytes("PreDefinition", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15() { _start: { lean_object* x_1; @@ -7522,17 +7781,17 @@ x_1 = lean_mk_string_from_bytes("Eqns", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17() { _start: { lean_object* x_1; @@ -7540,31 +7799,31 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17; +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19() { +static lean_object* _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18; -x_2 = lean_unsigned_to_nat(2381u); +x_1 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18; +x_2 = lean_unsigned_to_nat(2383u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1; +x_2 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1; x_3 = l_Lean_Meta_registerGetEqnsFn(x_2, x_1); if (lean_obj_tag(x_3) == 0) { @@ -7572,7 +7831,7 @@ lean_object* x_4; lean_object* x_5; lean_object* x_6; x_4 = lean_ctor_get(x_3, 1); lean_inc(x_4); lean_dec(x_3); -x_5 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2; +x_5 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2; x_6 = l_Lean_Meta_registerGetUnfoldEqnFn(x_5, x_4); if (lean_obj_tag(x_6) == 0) { @@ -7582,7 +7841,7 @@ lean_inc(x_7); lean_dec(x_6); x_8 = l___private_Lean_Elab_PreDefinition_WF_Eqns_0__Lean_Elab_WF_getFixedPrefix___lambda__2___closed__5; x_9 = 0; -x_10 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19; +x_10 = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19; x_11 = l_Lean_registerTraceClass(x_8, x_9, x_10, x_7); return x_11; } @@ -7819,15 +8078,15 @@ l_Lean_Elab_WF_mkEqns___closed__2 = _init_l_Lean_Elab_WF_mkEqns___closed__2(); lean_mark_persistent(l_Lean_Elab_WF_mkEqns___closed__2); l_Lean_Elab_WF_mkEqns___closed__3 = _init_l_Lean_Elab_WF_mkEqns___closed__3(); lean_mark_persistent(l_Lean_Elab_WF_mkEqns___closed__3); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__1); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__2); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__3); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111____closed__4); -if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2111_(lean_io_mk_world()); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__1); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__2); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__3); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__4(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113____closed__4); +if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2113_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_WF_eqnInfoExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_WF_eqnInfoExt); @@ -7864,45 +8123,45 @@ l_Lean_Elab_WF_registerEqnsInfo___closed__14 = _init_l_Lean_Elab_WF_registerEqns lean_mark_persistent(l_Lean_Elab_WF_registerEqnsInfo___closed__14); l_Lean_Elab_WF_registerEqnsInfo___closed__15 = _init_l_Lean_Elab_WF_registerEqnsInfo___closed__15(); lean_mark_persistent(l_Lean_Elab_WF_registerEqnsInfo___closed__15); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__1); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__2); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__3); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__4); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__5); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__6); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__7); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__8); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__9); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__10); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__11); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__12); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__13); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__14); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__15); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__16); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__17); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__18); -l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19(); -lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381____closed__19); -if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2381_(lean_io_mk_world()); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__1); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__2); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__3); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__4); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__5); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__6); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__7); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__8); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__9); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__10); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__11); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__12); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__13); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__14); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__15); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__16); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__17); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__18); +l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19 = _init_l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19(); +lean_mark_persistent(l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383____closed__19); +if (builtin) {res = l_Lean_Elab_WF_initFn____x40_Lean_Elab_PreDefinition_WF_Eqns___hyg_2383_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c index 4aea989eb666..42bbe7bf28e6 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/Main.c @@ -19,129 +19,145 @@ lean_object* l_Lean_Expr_bindingName_x21(lean_object*); lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_getFixedPrefix___spec__13___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda___at_Lean_Elab_getFixedPrefix___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_instInhabitedPreDefinition; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef(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__4(size_t, size_t, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1; +extern lean_object* l_Lean_allowUnsafeReducibility; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__4(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__1; lean_object* lean_mk_empty_array_with_capacity(lean_object*); lean_object* l_Lean_mkAppN(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16; static lean_object* l_Lean_Elab_wfRecursion___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__7(lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_getFixedPrefix___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppNumArgsAux(lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_varyingVarNames___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_Elab_wfRecursion___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_object*); -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_wfRecursion___spec__7(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__4; -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2; +static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15; lean_object* l_Lean_Meta_lambdaTelescope___at___private_Lean_Meta_Eqns_0__Lean_Meta_mkSimpleEqThm___spec__1___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__10___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14; lean_object* l_Lean_Meta_isExprDefEq(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__7; lean_object* l_Lean_Meta_markAsRecursive(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__4; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_sort___override(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_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4; +lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__1___boxed(lean_object**); +LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__9; lean_object* l_ST_Prim_Ref_get___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__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* lean_array_fget(lean_object*, lean_object*); +extern lean_object* l_Lean_Expr_instBEq; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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*); static lean_object* l_Lean_Elab_wfRecursion___lambda__2___closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___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*); static lean_object* l_Lean_Elab_varyingVarNames___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5; lean_object* l_Lean_Elab_addAndCompilePartialRec(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7; lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_Term_withAuxDecl___spec__1___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_withCommonTelescope_go___spec__1(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11(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_getFixedPrefix___spec__14___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_visitLet___at_Lean_Elab_getFixedPrefix___spec__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 l_Lean_Expr_isAppOf(lean_object*, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_withCommonTelescope_go___spec__2(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__2; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11; lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_MatcherApp_inferMatchType___spec__9(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15; +extern lean_object* l_Lean_maxRecDepth; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__3; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__6; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__5; static lean_object* l_panic___at_Lean_Elab_varyingVarNames___spec__1___closed__1; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__10; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withCommonTelescope_go___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_wfRecursion___lambda__2___closed__2; +lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___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* l_Lean_Meta_unfoldDeclsFrom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_visitLet_visit___at_Lean_Elab_getFixedPrefix___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_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11; lean_object* l_Lean_Meta_iteToDIte___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_getFixedPrefix___spec__13(lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5; +lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__16; lean_object* l_List_mapTR_loop___at_Lean_mkConstWithLevelParams___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___lambda__2___closed__5; +static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3; lean_object* l_Lean_HashMap_insert___at_Lean_ForEachExpr_visit___spec__3(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_instHashableArray___rarg___boxed(lean_object*, lean_object*); uint8_t l_Lean_Expr_isLambda(lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_wfRecursion___lambda__7___closed__1; lean_object* l_Lean_Meta_forallBoundedTelescope___at___private_Lean_PrettyPrinter_Delaborator_FieldNotation_0__Lean_PrettyPrinter_Delaborator_generalizedFieldInfo___spec__3___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_PreDefinition_WF_Main___hyg_2830____closed__4; size_t lean_usize_of_nat(lean_object*); lean_object* l_Lean_Elab_applyAttributesOf(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_instMonadControlTOfMonadControl___rarg(lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_preprocess___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__2; -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7; +static lean_object* l_Lean_Elab_wfRecursion___lambda__6___closed__2; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__3; -static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1; +static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4; lean_object* lean_st_ref_take(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*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef___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_PreDefinition_WF_Main___hyg_2773____closed__7; +static lean_object* l_Lean_Elab_wfRecursion___lambda__6___closed__1; lean_object* l_Lean_Meta_iteToDIte___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_withCommonTelescope___rarg___closed__1; -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__18; lean_object* l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3___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_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__14; +lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_mkFix(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitForall___at_Lean_Elab_getFixedPrefix___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* 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*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_Elab_addNonRec___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_wfRecursion___lambda__7___closed__2; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13; lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__2(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___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* l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__2___closed__1; @@ -153,22 +169,21 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_varyingVarNames___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__4(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); lean_object* l_ReaderT_instApplicativeOfMonad___rarg(lean_object*); -static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5; +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_withCommonTelescope___spec__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withCommonTelescope_go___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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* l_Lean_Name_num___override(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__5(lean_object*, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__10___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__7___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1; static lean_object* l_Lean_Elab_varyingVarNames___closed__8; lean_object* l_Lean_Elab_WF_TerminationArgument_elab(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3(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_wfRecursion___lambda__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_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ST_Prim_Ref_modifyGetUnsafe___rarg___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -176,33 +191,43 @@ static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinitio lean_object* l_Lean_Elab_abstractNestedProofs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__2___closed__2; extern lean_object* l_Lean_instInhabitedExpr; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_ForEachExpr_visit___spec__1(lean_object*, lean_object*); +lean_object* l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(lean_object*, lean_object*, uint8_t); lean_object* l_Lean_Elab_WF_registerEqnsInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___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_getFixedPrefix___spec__14___closed__2; +lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_getFixedPrefix___spec__6___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__8; +lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); +static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_withCommonTelescope(lean_object*); lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_getFixedPrefix___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_getFixedPrefix___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Elab_WF_preprocess___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__6; 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773_(lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(size_t, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_wfRecursion___lambda__5___closed__1; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6; +extern lean_object* l_Lean_diagnostics; static lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3___closed__2; +lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___closed__3; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10; LEAN_EXPORT lean_object* l_Lean_Meta_visitLet_visit___at_Lean_Elab_getFixedPrefix___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*, lean_object*, 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*); 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*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__4; static lean_object* l_Lean_Elab_wfRecursion___lambda__4___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__6(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_getFixedPrefix___spec__13___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__3___boxed(lean_object**); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__12; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_withCommonTelescope_go___spec__3___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_instInhabitedMetaM___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -213,18 +238,20 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_withCommonTelesco LEAN_EXPORT lean_object* l_Lean_Meta_withLetDecl___at_Lean_Elab_getFixedPrefix___spec__13___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*); uint8_t l_Lean_Expr_binderInfo(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_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_getFixedPrefix___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_withCommonTelescope___spec__1(size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withCommonTelescope___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_wfRecursion___lambda__1___closed__1; lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_addAndCompileUnsafe___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_Lean_Elab_wfRecursion___spec__11(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3; +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___boxed(lean_object**); lean_object* l_List_mapTR_loop___at_Lean_Elab_addAndCompileUnsafe___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19; lean_object* l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1; -static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2; +static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1; lean_object* l_Lean_Elab_WF_guessLex(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_bindingDomain_x21(lean_object*); @@ -235,80 +262,100 @@ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix 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_wfRecursion___lambda__7___closed__3; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4; lean_object* l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__15; lean_object* l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_wfRecursion___closed__3; +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17; uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); static lean_object* l_Lean_Elab_varyingVarNames___closed__7; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5; +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___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_object*, lean_object*); static lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3___closed__1; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8; static lean_object* l_Lean_Elab_wfRecursion___closed__2; lean_object* l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_forallBoundedTelescope___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_addAsAxiom(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__6; +extern lean_object* l_Lean_Expr_instHashable; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_getFixedPrefix___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_nat_sub(lean_object*, lean_object*); lean_object* l_instMonadControlReaderT___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_getFixedPrefix___spec__1(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__10(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___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_wfRecursion___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_Lean_Elab_WF_packMutual(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_varyingVarNames___closed__4; +lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Array_ofSubarray___rarg(lean_object*); static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27_visit___at_Lean_Elab_getFixedPrefix___spec__4___lambda__1(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18; extern lean_object* l_Lean_Core_instMonadCoreM; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___boxed(lean_object**); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_withCommonTelescope_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*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_withCommonTelescope_go___spec__3(lean_object*, size_t, size_t); -static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__3; static lean_object* l_Lean_Elab_wfRecursion___lambda__1___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_expr_instantiate_rev(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_varyingVarNames___lambda__1___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_PreDefinition_WF_Main___hyg_2830____closed__8; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); lean_object* l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___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_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__9(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__7(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_varyingVarNames___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_varyingVarNames(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_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__1(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1; static lean_object* l_Lean_Elab_wfRecursion___closed__1; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_visitForall_visit___at_Lean_Elab_getFixedPrefix___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__1; lean_object* l_Lean_Expr_bindingBody_x21(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__8; lean_object* l___private_Lean_Expr_0__Lean_Expr_getAppArgsAux(lean_object*, lean_object*, lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_forEachExpr_x27___at_Lean_Elab_getFixedPrefix___spec__3___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_varyingVarNames___closed__6; lean_object* lean_string_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17; +lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__7___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___boxed(lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Elab_wfRecursion___lambda__7___closed__4; -static lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4; lean_object* l_Lean_Meta_mkLambdaFVars(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17; +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_nat_add(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5; -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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_wfRecursion___lambda__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef___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_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___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_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_wfRecursion___spec__9(lean_object*); uint8_t l_Lean_Expr_isForall(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withCommonTelescope_go(lean_object*); +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__13; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3; +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13(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_getFixedPrefix___lambda__2___closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -319,11 +366,13 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spe LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Elab_getFixedPrefix___spec__7___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_getFixedPrefix___spec__14___lambda__3___closed__6; 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*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__16; +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4(lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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* l_ReaderT_instMonad___rarg(lean_object*); +static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__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*); LEAN_EXPORT lean_object* l_Lean_Meta_visitLambda_visit___at_Lean_Elab_getFixedPrefix___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__10; +static lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_varyingVarNames___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); lean_object* l_instMonadControlReaderT___lambda__3___boxed(lean_object*, lean_object*, lean_object*); @@ -6234,1356 +6283,1915 @@ return x_26; } } } -static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1() { +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__1() { +_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_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = lean_unsigned_to_nat(1u); -x_2 = lean_mk_empty_array_with_capacity(x_1); +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___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_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3() { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -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_6); -lean_dec(x_5); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_11); -return x_13; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; +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; } -else -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; -lean_dec(x_4); -x_14 = lean_array_uget(x_1, x_3); -x_15 = lean_ctor_get(x_14, 3); -lean_inc(x_15); -x_16 = l_Lean_Meta_markAsRecursive(x_15, x_9, x_10, x_11); -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1; -x_19 = lean_array_push(x_18, x_14); -x_20 = 1; -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_21 = l_Lean_Elab_applyAttributesOf(x_19, x_20, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -lean_dec(x_19); -if (lean_obj_tag(x_21) == 0) +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4() { +_start: { -lean_object* x_22; size_t x_23; size_t x_24; lean_object* x_25; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -x_23 = 1; -x_24 = lean_usize_add(x_3, x_23); -x_25 = lean_box(0); -x_3 = x_24; -x_4 = x_25; -x_11 = x_22; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3; +x_2 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5() { +_start: { -uint8_t x_27; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_27 = !lean_is_exclusive(x_21); -if (x_27 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; +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_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__6() { +_start: { -return x_21; +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_InfoCacheKey_instHashable___boxed), 1, 0); +return x_1; } -else +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7() { +_start: { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_21, 0); -x_29 = lean_ctor_get(x_21, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_21); -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_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; +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_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instBEqLocalInstance___boxed), 2, 0); +return x_1; } } +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8; +x_2 = lean_alloc_closure((void*)(l_Array_instBEq___rarg___boxed), 3, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__4(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) { +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__10() { _start: { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_10 = lean_st_ref_get(x_8, x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_ctor_get(x_11, 6); -lean_inc(x_13); -lean_dec(x_11); -x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); -lean_dec(x_13); -x_15 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -lean_inc(x_3); -x_17 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16); -if (lean_obj_tag(x_17) == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9; +x_2 = l_Lean_Expr_instBEq; +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_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11() { +_start: { -lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_instHashableLocalInstance___boxed), 1, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12() { +_start: { -lean_object* x_22; -x_22 = lean_ctor_get(x_20, 0); -lean_dec(x_22); -lean_ctor_set(x_20, 0, x_18); -return x_20; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11; +x_2 = lean_alloc_closure((void*)(l_instHashableArray___rarg___boxed), 2, 1); +lean_closure_set(x_2, 0, x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__13() { +_start: { -lean_object* x_23; lean_object* x_24; -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_23); -lean_dec(x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_18); -lean_ctor_set(x_24, 1, x_23); -return x_24; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12; +x_2 = l_Lean_Expr_instHashable; +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; } } -else +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14() { +_start: { -lean_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_ctor_get(x_17, 0); -lean_inc(x_25); -x_26 = lean_ctor_get(x_17, 1); -lean_inc(x_26); -lean_dec(x_17); -x_27 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_14, x_3, x_4, x_5, x_6, x_7, x_8, 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); -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; +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_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__15() { +_start: { -lean_object* x_29; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -lean_ctor_set_tag(x_27, 1); -lean_ctor_set(x_27, 0, x_25); -return x_27; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Expr_instBEq; +x_2 = lean_alloc_closure((void*)(l_instBEqProd___rarg), 4, 2); +lean_closure_set(x_2, 0, x_1); +lean_closure_set(x_2, 1, x_1); +return x_2; } -else +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__16() { +_start: { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_31, 0, x_25); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Expr_instHashable; +x_2 = lean_alloc_closure((void*)(l_instHashableProd___rarg___boxed), 3, 2); +lean_closure_set(x_2, 0, x_1); +lean_closure_set(x_2, 1, x_1); +return x_2; } } +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2; +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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18() { _start: { -uint8_t x_5; -x_5 = lean_usize_dec_lt(x_3, x_2); -if (x_5 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17; +x_2 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_2, 0, x_1); +lean_ctor_set(x_2, 1, x_1); +lean_ctor_set(x_2, 2, x_1); +lean_ctor_set(x_2, 3, x_1); +return x_2; +} +} +static lean_object* _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19() { +_start: { -return x_4; +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; +x_1 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5; +x_2 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7; +x_3 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14; +x_4 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18; +x_5 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_5, 0, x_1); +lean_ctor_set(x_5, 1, x_2); +lean_ctor_set(x_5, 2, x_3); +lean_ctor_set(x_5, 3, x_1); +lean_ctor_set(x_5, 4, x_1); +lean_ctor_set(x_5, 5, x_4); +lean_ctor_set(x_5, 6, x_4); +return x_5; } -else +} +LEAN_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__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, lean_object* x_9) { +_start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_array_uget(x_1, x_3); -x_7 = lean_ctor_get(x_6, 6); -lean_inc(x_7); -x_8 = lean_ctor_get(x_7, 2); -lean_inc(x_8); -lean_dec(x_7); -if (lean_obj_tag(x_8) == 0) +lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; +x_10 = lean_st_ref_take(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_is_exclusive(x_11); +if (x_13 == 0) { -uint8_t x_9; -x_9 = !lean_is_exclusive(x_4); -if (x_9 == 0) +lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; 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_14 = lean_ctor_get(x_11, 0); +x_15 = lean_ctor_get(x_11, 4); +lean_dec(x_15); +x_16 = 0; +x_17 = lean_box(0); +x_18 = l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(x_14, x_1, x_2, x_16, x_17); +x_19 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4; +lean_ctor_set(x_11, 4, x_19); +lean_ctor_set(x_11, 0, x_18); +x_20 = lean_st_ref_set(x_8, x_11, x_12); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +x_22 = lean_st_ref_take(x_6, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); +lean_dec(x_22); +x_25 = !lean_is_exclusive(x_23); +if (x_25 == 0) { -lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; -x_10 = lean_ctor_get(x_4, 1); -x_11 = lean_array_push(x_10, x_6); -lean_ctor_set(x_4, 1, x_11); -x_12 = 1; -x_13 = lean_usize_add(x_3, x_12); -x_3 = x_13; -goto _start; +lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; +x_26 = lean_ctor_get(x_23, 1); +lean_dec(x_26); +x_27 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19; +lean_ctor_set(x_23, 1, x_27); +x_28 = lean_st_ref_set(x_6, x_23, x_24); +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_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; -x_15 = lean_ctor_get(x_4, 0); -x_16 = lean_ctor_get(x_4, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_4); -x_17 = lean_array_push(x_16, x_6); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_15); -lean_ctor_set(x_18, 1, x_17); -x_19 = 1; -x_20 = lean_usize_add(x_3, x_19); -x_3 = x_20; -x_4 = x_18; -goto _start; +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 { -uint8_t x_22; -lean_dec(x_8); -x_22 = !lean_is_exclusive(x_4); -if (x_22 == 0) -{ -lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; -x_23 = lean_ctor_get(x_4, 0); -x_24 = lean_array_push(x_23, x_6); -lean_ctor_set(x_4, 0, x_24); -x_25 = 1; -x_26 = lean_usize_add(x_3, x_25); -x_3 = x_26; -goto _start; +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_35 = lean_ctor_get(x_23, 0); +x_36 = lean_ctor_get(x_23, 2); +x_37 = lean_ctor_get(x_23, 3); +x_38 = lean_ctor_get(x_23, 4); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_23); +x_39 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19; +x_40 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_40, 0, x_35); +lean_ctor_set(x_40, 1, x_39); +lean_ctor_set(x_40, 2, x_36); +lean_ctor_set(x_40, 3, x_37); +lean_ctor_set(x_40, 4, x_38); +x_41 = lean_st_ref_set(x_6, x_40, x_24); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_43 = x_41; +} else { + lean_dec_ref(x_41); + x_43 = lean_box(0); } -else -{ -lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; -x_28 = lean_ctor_get(x_4, 0); -x_29 = lean_ctor_get(x_4, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_4); -x_30 = lean_array_push(x_28, x_6); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -x_32 = 1; -x_33 = lean_usize_add(x_3, x_32); -x_3 = x_33; -x_4 = x_31; -goto _start; +x_44 = lean_box(0); +if (lean_is_scalar(x_43)) { + x_45 = lean_alloc_ctor(0, 2, 0); +} else { + x_45 = x_43; +} +lean_ctor_set(x_45, 0, x_44); +lean_ctor_set(x_45, 1, x_42); +return x_45; } } +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_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_46 = lean_ctor_get(x_11, 0); +x_47 = lean_ctor_get(x_11, 1); +x_48 = lean_ctor_get(x_11, 2); +x_49 = lean_ctor_get(x_11, 3); +x_50 = lean_ctor_get(x_11, 5); +x_51 = lean_ctor_get(x_11, 6); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_11); +x_52 = 0; +x_53 = lean_box(0); +x_54 = l___private_Lean_ReducibilityAttrs_0__Lean_setReducibilityStatusCore(x_46, x_1, x_2, x_52, x_53); +x_55 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4; +x_56 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_47); +lean_ctor_set(x_56, 2, x_48); +lean_ctor_set(x_56, 3, x_49); +lean_ctor_set(x_56, 4, x_55); +lean_ctor_set(x_56, 5, x_50); +lean_ctor_set(x_56, 6, x_51); +x_57 = lean_st_ref_set(x_8, x_56, x_12); +x_58 = lean_ctor_get(x_57, 1); +lean_inc(x_58); +lean_dec(x_57); +x_59 = lean_st_ref_take(x_6, x_58); +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_ctor_get(x_59, 1); +lean_inc(x_61); +lean_dec(x_59); +x_62 = lean_ctor_get(x_60, 0); +lean_inc(x_62); +x_63 = lean_ctor_get(x_60, 2); +lean_inc(x_63); +x_64 = lean_ctor_get(x_60, 3); +lean_inc(x_64); +x_65 = lean_ctor_get(x_60, 4); +lean_inc(x_65); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + lean_ctor_release(x_60, 2); + lean_ctor_release(x_60, 3); + lean_ctor_release(x_60, 4); + x_66 = x_60; +} else { + lean_dec_ref(x_60); + x_66 = lean_box(0); } +x_67 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19; +if (lean_is_scalar(x_66)) { + x_68 = lean_alloc_ctor(0, 5, 0); +} else { + x_68 = x_66; } +lean_ctor_set(x_68, 0, x_62); +lean_ctor_set(x_68, 1, x_67); +lean_ctor_set(x_68, 2, x_63); +lean_ctor_set(x_68, 3, x_64); +lean_ctor_set(x_68, 4, x_65); +x_69 = lean_st_ref_set(x_6, x_68, x_61); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_71 = x_69; +} else { + lean_dec_ref(x_69); + x_71 = lean_box(0); } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Missing `termination_by`; this function is mutually ", 52); -return x_1; +x_72 = lean_box(0); +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; } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("recursive with ", 15); +x_1 = lean_mk_string_from_bytes("semireducible", 13); return x_1; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes(", which has a `termination_by` clause.", 38); +x_1 = lean_mk_string_from_bytes("reducible", 9); return x_1; } } -static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6() { +static lean_object* _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = lean_box(0); +x_2 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3; +x_3 = l_Lean_Name_str___override(x_1, x_2); +return x_3; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { +LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4(lean_object* x_1, size_t x_2, size_t x_3) { _start: { -uint8_t x_13; -x_13 = lean_usize_dec_eq(x_3, x_4); -if (x_13 == 0) +uint8_t x_4; +x_4 = lean_usize_dec_eq(x_2, x_3); +if (x_4 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_object* x_5; lean_object* x_6; lean_object* x_7; uint8_t x_8; +x_5 = lean_array_uget(x_1, x_2); +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); lean_dec(x_5); -x_14 = lean_array_uget(x_2, x_3); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_array_get_size(x_1); -x_17 = lean_unsigned_to_nat(0u); -x_18 = lean_nat_dec_lt(x_17, x_16); -lean_dec(x_16); -if (x_18 == 0) +x_7 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2; +x_8 = lean_name_eq(x_6, x_7); +if (x_8 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; size_t x_33; size_t x_34; -x_19 = l_Lean_Elab_instInhabitedPreDefinition; -x_20 = l___private_Init_GetElem_0__outOfBounds___rarg(x_19); -x_21 = lean_ctor_get(x_20, 3); -lean_inc(x_21); -lean_dec(x_20); -x_22 = l_Lean_MessageData_ofName(x_21); -x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___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_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6; -x_26 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -x_27 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___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 = 2; -x_30 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_15, x_28, x_29, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_15); -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 = 1; -x_34 = lean_usize_add(x_3, x_33); -x_3 = x_34; -x_5 = x_31; -x_12 = x_32; +lean_object* x_9; uint8_t x_10; +x_9 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4; +x_10 = lean_name_eq(x_6, x_9); +lean_dec(x_6); +if (x_10 == 0) +{ +size_t x_11; size_t x_12; +x_11 = 1; +x_12 = lean_usize_add(x_2, x_11); +x_2 = x_12; goto _start; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; -x_36 = lean_array_fget(x_1, x_17); -x_37 = lean_ctor_get(x_36, 3); -lean_inc(x_37); -lean_dec(x_36); -x_38 = l_Lean_MessageData_ofName(x_37); -x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4; -x_40 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_40, 1, x_38); -x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_43); -lean_ctor_set(x_44, 1, x_42); -x_45 = 2; -x_46 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_15, x_44, x_45, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -lean_dec(x_15); -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 = 1; -x_50 = lean_usize_add(x_3, x_49); -x_3 = x_50; -x_5 = x_47; -x_12 = x_48; -goto _start; +uint8_t x_14; +x_14 = 1; +return x_14; } } else { -lean_object* x_52; -x_52 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_52, 0, x_5); -lean_ctor_set(x_52, 1, x_12); -return x_52; -} +uint8_t x_15; +lean_dec(x_6); +x_15 = 1; +return x_15; } } -LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_wfRecursion___spec__7(lean_object* x_1) { -_start: +else { -lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_WF_instInhabitedTerminationBy; -x_3 = lean_panic_fn(x_2, x_1); -return x_3; -} +uint8_t x_16; +x_16 = 0; +return x_16; } -static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); -return x_1; } } -static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2() { +static lean_object* _init_l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1() { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Option.get!", 11); -return x_1; +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_mk_empty_array_with_capacity(x_1); +return x_2; } } -static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3() { +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(size_t x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("value is none", 13); -return x_1; -} -} -static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4() { -_start: +uint8_t x_13; +x_13 = lean_usize_dec_lt(x_4, x_3); +if (x_13 == 0) { -lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; -x_1 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1; -x_2 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2; -x_3 = lean_unsigned_to_nat(16u); -x_4 = lean_unsigned_to_nat(14u); -x_5 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3; -x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); -return x_6; -} +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_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_5); +lean_ctor_set(x_14, 1, x_12); +return x_14; } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { -_start: -{ -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) +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_18 = lean_unsigned_to_nat(1u); -x_19 = lean_nat_sub(x_5, x_18); +lean_object* x_15; lean_object* x_16; lean_object* x_17; 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_dec(x_5); -x_20 = lean_array_fget(x_4, x_6); -x_21 = lean_array_get_size(x_2); -x_22 = lean_nat_dec_lt(x_6, 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; -x_23 = lean_ctor_get(x_20, 6); +x_15 = lean_array_uget(x_2, x_4); +x_23 = lean_ctor_get(x_15, 3); lean_inc(x_23); -x_24 = lean_ctor_get(x_20, 3); -lean_inc(x_24); -x_25 = lean_ctor_get(x_20, 4); +lean_inc(x_23); +x_24 = l_Lean_Meta_markAsRecursive(x_23, x_10, x_11, x_12); +x_25 = lean_ctor_get(x_24, 1); lean_inc(x_25); -lean_dec(x_20); -x_26 = lean_ctor_get(x_23, 4); -lean_inc(x_26); -x_27 = lean_ctor_get(x_23, 2); -lean_inc(x_27); -lean_dec(x_23); -x_28 = l_Lean_Elab_withCommonTelescope___rarg___closed__1; -x_29 = l___private_Init_GetElem_0__outOfBounds___rarg(x_28); -x_30 = lean_array_get_size(x_29); -lean_dec(x_29); -x_31 = lean_nat_add(x_1, x_30); -lean_dec(x_30); -if (lean_obj_tag(x_27) == 0) -{ -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4; -x_33 = l_panic___at_Lean_Elab_wfRecursion___spec__7(x_32); -lean_inc(x_14); -lean_inc(x_13); -lean_inc(x_12); +lean_dec(x_24); +x_26 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1; +lean_inc(x_15); +x_27 = lean_array_push(x_26, x_15); +x_28 = 1; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); -x_34 = l_Lean_Elab_WF_TerminationArgument_elab(x_24, x_25, x_31, x_26, x_33, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_34) == 0) +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_29 = l_Lean_Elab_applyAttributesOf(x_27, x_28, x_6, x_7, x_8, x_9, x_10, x_11, x_25); +lean_dec(x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_ctor_get(x_34, 1); -lean_inc(x_36); -lean_dec(x_34); -x_37 = lean_nat_add(x_6, x_18); -lean_dec(x_6); -x_38 = lean_array_push(x_8, x_35); -x_5 = x_19; -x_6 = x_37; -x_7 = lean_box(0); -x_8 = x_38; -x_15 = x_36; -goto _start; +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_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = lean_ctor_get(x_15, 2); +lean_inc(x_31); +lean_dec(x_15); +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); +x_33 = lean_array_get_size(x_32); +x_34 = lean_unsigned_to_nat(0u); +x_35 = lean_nat_dec_lt(x_34, x_33); +if (x_35 == 0) +{ +uint8_t x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +lean_dec(x_33); +lean_dec(x_32); +x_36 = 2; +x_37 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3(x_23, x_36, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___closed__1; +x_16 = x_39; +x_17 = x_38; +goto block_22; } else { -uint8_t x_40; -lean_dec(x_19); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); +size_t x_40; uint8_t x_41; +x_40 = lean_usize_of_nat(x_33); +lean_dec(x_33); +x_41 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4(x_32, x_1, x_40); +lean_dec(x_32); +if (x_41 == 0) +{ +uint8_t x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_42 = 2; +x_43 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3(x_23, x_42, x_6, x_7, x_8, x_9, x_10, x_11, x_30); +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___closed__1; +x_16 = x_45; +x_17 = x_44; +goto block_22; +} +else +{ +lean_object* x_46; +lean_dec(x_23); +x_46 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___lambda__2___closed__1; +x_16 = x_46; +x_17 = x_30; +goto block_22; +} +} +} +else +{ +uint8_t x_47; +lean_dec(x_23); +lean_dec(x_15); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_40 = !lean_is_exclusive(x_34); -if (x_40 == 0) +x_47 = !lean_is_exclusive(x_29); +if (x_47 == 0) { -return x_34; +return x_29; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_34, 0); -x_42 = lean_ctor_get(x_34, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_34); -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_48; lean_object* x_49; lean_object* x_50; +x_48 = lean_ctor_get(x_29, 0); +x_49 = lean_ctor_get(x_29, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_29); +x_50 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_50, 0, x_48); +lean_ctor_set(x_50, 1, x_49); +return x_50; } } -else +block_22: { -lean_object* x_44; lean_object* x_45; -x_44 = lean_ctor_get(x_27, 0); -lean_inc(x_44); -lean_dec(x_27); -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_45 = l_Lean_Elab_WF_TerminationArgument_elab(x_24, x_25, x_31, x_26, x_44, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_45) == 0) -{ -lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_46 = 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_nat_add(x_6, x_18); -lean_dec(x_6); -x_49 = lean_array_push(x_8, x_46); -x_5 = x_19; -x_6 = x_48; -x_7 = lean_box(0); -x_8 = x_49; -x_15 = x_47; +lean_object* x_18; size_t x_19; size_t x_20; +x_18 = lean_ctor_get(x_16, 0); +lean_inc(x_18); +lean_dec(x_16); +x_19 = 1; +x_20 = lean_usize_add(x_4, x_19); +x_4 = x_20; +x_5 = x_18; +x_12 = x_17; goto _start; } -else +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6(uint8_t x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: { -uint8_t x_51; -lean_dec(x_19); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(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; lean_object* x_17; +x_10 = lean_st_ref_get(x_8, x_9); +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_10, 1); +lean_inc(x_12); lean_dec(x_10); -lean_dec(x_9); +x_13 = lean_ctor_get(x_11, 6); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_ctor_get_uint8(x_13, sizeof(void*)*2); +lean_dec(x_13); +x_15 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_17 = lean_apply_7(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_16); +if (lean_obj_tag(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_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_19); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_51 = !lean_is_exclusive(x_45); -if (x_51 == 0) +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_21 = !lean_is_exclusive(x_20); +if (x_21 == 0) { -return x_45; +lean_object* x_22; +x_22 = lean_ctor_get(x_20, 0); +lean_dec(x_22); +lean_ctor_set(x_20, 0, x_18); +return x_20; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_45, 0); -x_53 = lean_ctor_get(x_45, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_45); -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_23; lean_object* x_24; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } 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; -x_55 = lean_ctor_get(x_20, 6); -lean_inc(x_55); -x_56 = lean_ctor_get(x_20, 3); -lean_inc(x_56); -x_57 = lean_ctor_get(x_20, 4); -lean_inc(x_57); -lean_dec(x_20); -x_58 = lean_ctor_get(x_55, 4); -lean_inc(x_58); -x_59 = lean_ctor_get(x_55, 2); -lean_inc(x_59); -lean_dec(x_55); -x_60 = lean_array_fget(x_2, x_6); -x_61 = lean_array_get_size(x_60); -lean_dec(x_60); -x_62 = lean_nat_add(x_1, x_61); -lean_dec(x_61); -if (lean_obj_tag(x_59) == 0) -{ -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4; -x_64 = l_panic___at_Lean_Elab_wfRecursion___spec__7(x_63); -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_65 = l_Lean_Elab_WF_TerminationArgument_elab(x_56, x_57, x_62, x_58, x_64, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_65) == 0) -{ -lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; -x_66 = lean_ctor_get(x_65, 0); -lean_inc(x_66); -x_67 = lean_ctor_get(x_65, 1); -lean_inc(x_67); -lean_dec(x_65); -x_68 = lean_nat_add(x_6, x_18); -lean_dec(x_6); -x_69 = lean_array_push(x_8, x_66); -x_5 = x_19; -x_6 = x_68; -x_7 = lean_box(0); -x_8 = x_69; -x_15 = x_67; -goto _start; -} -else -{ -uint8_t x_71; -lean_dec(x_19); -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_object* x_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_25 = lean_ctor_get(x_17, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); +lean_dec(x_17); +x_27 = l_Lean_Elab_enableInfoTree___at_Lean_Elab_addAndCompilePartialRec___spec__5(x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_26); lean_dec(x_8); +lean_dec(x_7); lean_dec(x_6); -x_71 = !lean_is_exclusive(x_65); -if (x_71 == 0) +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_28 = !lean_is_exclusive(x_27); +if (x_28 == 0) { -return x_65; +lean_object* x_29; +x_29 = lean_ctor_get(x_27, 0); +lean_dec(x_29); +lean_ctor_set_tag(x_27, 1); +lean_ctor_set(x_27, 0, x_25); +return x_27; } else { -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_65, 0); -x_73 = lean_ctor_get(x_65, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_65); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_72); -lean_ctor_set(x_74, 1, x_73); -return x_74; +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_27, 1); +lean_inc(x_30); +lean_dec(x_27); +x_31 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_31, 0, x_25); +lean_ctor_set(x_31, 1, x_30); +return x_31; +} } } } +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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) +{ +return x_4; +} else { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_59, 0); -lean_inc(x_75); -lean_dec(x_59); -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_76 = l_Lean_Elab_WF_TerminationArgument_elab(x_56, x_57, x_62, x_58, x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15); -if (lean_obj_tag(x_76) == 0) +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_array_uget(x_1, x_3); +x_7 = lean_ctor_get(x_6, 6); +lean_inc(x_7); +x_8 = lean_ctor_get(x_7, 2); +lean_inc(x_8); +lean_dec(x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -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_nat_add(x_6, x_18); -lean_dec(x_6); -x_80 = lean_array_push(x_8, x_77); -x_5 = x_19; -x_6 = x_79; -x_7 = lean_box(0); -x_8 = x_80; -x_15 = x_78; +uint8_t x_9; +x_9 = !lean_is_exclusive(x_4); +if (x_9 == 0) +{ +lean_object* x_10; lean_object* x_11; size_t x_12; size_t x_13; +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_array_push(x_10, x_6); +lean_ctor_set(x_4, 1, x_11); +x_12 = 1; +x_13 = lean_usize_add(x_3, x_12); +x_3 = x_13; goto _start; } else { -uint8_t x_82; -lean_dec(x_19); -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_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; size_t x_19; size_t x_20; +x_15 = lean_ctor_get(x_4, 0); +x_16 = lean_ctor_get(x_4, 1); +lean_inc(x_16); +lean_inc(x_15); +lean_dec(x_4); +x_17 = lean_array_push(x_16, x_6); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_15); +lean_ctor_set(x_18, 1, x_17); +x_19 = 1; +x_20 = lean_usize_add(x_3, x_19); +x_3 = x_20; +x_4 = x_18; +goto _start; +} +} +else +{ +uint8_t x_22; lean_dec(x_8); -lean_dec(x_6); -x_82 = !lean_is_exclusive(x_76); -if (x_82 == 0) +x_22 = !lean_is_exclusive(x_4); +if (x_22 == 0) { -return x_76; +lean_object* x_23; lean_object* x_24; size_t x_25; size_t x_26; +x_23 = lean_ctor_get(x_4, 0); +x_24 = lean_array_push(x_23, x_6); +lean_ctor_set(x_4, 0, x_24); +x_25 = 1; +x_26 = lean_usize_add(x_3, x_25); +x_3 = x_26; +goto _start; } else { -lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_83 = lean_ctor_get(x_76, 0); -x_84 = lean_ctor_get(x_76, 1); -lean_inc(x_84); -lean_inc(x_83); -lean_dec(x_76); -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_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; size_t x_32; size_t x_33; +x_28 = lean_ctor_get(x_4, 0); +x_29 = lean_ctor_get(x_4, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_4); +x_30 = lean_array_push(x_28, x_6); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +x_32 = 1; +x_33 = lean_usize_add(x_3, x_32); +x_3 = x_33; +x_4 = x_31; +goto _start; } } } } } -else +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1() { +_start: { -lean_object* x_86; -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_6); -lean_dec(x_5); -x_86 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_86, 0, x_8); -lean_ctor_set(x_86, 1, x_15); -return x_86; -} +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Missing `termination_by`; this function is mutually ", 52); +return x_1; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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) { +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2() { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -if (x_12 == 0) +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3() { +_start: { -lean_object* x_13; -lean_dec(x_10); -lean_dec(x_9); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_11); -return x_13; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("recursive with ", 15); +return x_1; } -else +} +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4() { +_start: { -lean_object* x_14; lean_object* x_15; -lean_dec(x_4); -x_14 = lean_array_uget(x_1, x_3); -lean_inc(x_10); -lean_inc(x_9); -x_15 = l_Lean_Elab_addAsAxiom(x_14, x_7, x_8, x_9, x_10, x_11); -lean_dec(x_14); -if (lean_obj_tag(x_15) == 0) -{ -lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = 1; -x_18 = lean_usize_add(x_3, x_17); -x_19 = lean_box(0); -x_3 = x_18; -x_4 = x_19; -x_11 = x_16; -goto _start; +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } -else -{ -uint8_t x_21; -lean_dec(x_10); -lean_dec(x_9); -x_21 = !lean_is_exclusive(x_15); -if (x_21 == 0) -{ -return x_15; } -else +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5() { +_start: { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_15, 0); -x_23 = lean_ctor_get(x_15, 1); -lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_15); -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_1; +x_1 = lean_mk_string_from_bytes(", which has a `termination_by` clause.", 38); +return x_1; } } +static lean_object* _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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, lean_object* x_10, lean_object* x_11) { +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -uint8_t x_12; -x_12 = lean_usize_dec_lt(x_3, x_2); -if (x_12 == 0) -{ -lean_object* x_13; -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_1); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_4); -lean_ctor_set(x_13, 1, x_11); -return x_13; -} -else +uint8_t x_13; +x_13 = lean_usize_dec_eq(x_3, x_4); +if (x_13 == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_14 = lean_array_uget(x_4, x_3); -x_15 = lean_unsigned_to_nat(0u); -x_16 = lean_array_uset(x_4, x_3, x_15); -lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_1); -x_17 = l_Lean_Elab_varyingVarNames(x_1, x_14, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_17) == 0) +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_5); +x_14 = lean_array_uget(x_2, x_3); +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +lean_dec(x_14); +x_16 = lean_array_get_size(x_1); +x_17 = lean_unsigned_to_nat(0u); +x_18 = lean_nat_dec_lt(x_17, x_16); +lean_dec(x_16); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; -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 = 1; -x_21 = lean_usize_add(x_3, x_20); -x_22 = lean_array_uset(x_16, x_3, x_18); -x_3 = x_21; -x_4 = x_22; -x_11 = x_19; +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; 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; size_t x_33; size_t x_34; +x_19 = l_Lean_Elab_instInhabitedPreDefinition; +x_20 = l___private_Init_GetElem_0__outOfBounds___rarg(x_19); +x_21 = lean_ctor_get(x_20, 3); +lean_inc(x_21); +lean_dec(x_20); +x_22 = l_Lean_MessageData_ofName(x_21); +x_23 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___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_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6; +x_26 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___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 = 2; +x_30 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_15, x_28, x_29, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_15); +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 = 1; +x_34 = lean_usize_add(x_3, x_33); +x_3 = x_34; +x_5 = x_31; +x_12 = x_32; goto _start; } else { -uint8_t x_24; -lean_dec(x_16); -lean_dec(x_10); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_1); -x_24 = !lean_is_exclusive(x_17); -if (x_24 == 0) -{ -return x_17; +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; size_t x_49; size_t x_50; +x_36 = lean_array_fget(x_1, x_17); +x_37 = lean_ctor_get(x_36, 3); +lean_inc(x_37); +lean_dec(x_36); +x_38 = l_Lean_MessageData_ofName(x_37); +x_39 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4; +x_40 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_40, 0, x_39); +lean_ctor_set(x_40, 1, x_38); +x_41 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6; +x_42 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_42, 0, x_40); +lean_ctor_set(x_42, 1, x_41); +x_43 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2; +x_44 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_44, 0, x_43); +lean_ctor_set(x_44, 1, x_42); +x_45 = 2; +x_46 = l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__2(x_15, x_44, x_45, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_15); +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 = 1; +x_50 = lean_usize_add(x_3, x_49); +x_3 = x_50; +x_5 = x_47; +x_12 = x_48; +goto _start; +} } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_17, 0); -x_26 = lean_ctor_get(x_17, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_17); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_object* x_52; +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_5); +lean_ctor_set(x_52, 1, x_12); +return x_52; } } } +LEAN_EXPORT lean_object* l_panic___at_Lean_Elab_wfRecursion___spec__9(lean_object* x_1) { +_start: +{ +lean_object* x_2; lean_object* x_3; +x_2 = l_Lean_Elab_WF_instInhabitedTerminationBy; +x_3 = lean_panic_fn(x_2, x_1); +return x_3; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1() { +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_iteToDIte___lambda__2___boxed), 6, 0); +x_1 = lean_mk_string_from_bytes("Init.Data.Option.BasicAux", 25); return x_1; } } -static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2() { +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_iteToDIte___lambda__1___boxed), 6, 0); +x_1 = lean_mk_string_from_bytes("Option.get!", 11); return x_1; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11(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) { +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3() { _start: { -uint8_t x_11; -x_11 = lean_usize_dec_lt(x_2, x_1); -if (x_11 == 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_3); -lean_ctor_set(x_12, 1, x_10); -return x_12; +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("value is none", 13); +return x_1; } -else +} +static lean_object* _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4() { +_start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_13 = lean_array_uget(x_3, x_2); -x_14 = lean_unsigned_to_nat(0u); -x_15 = lean_array_uset(x_3, x_2, x_14); -x_16 = !lean_is_exclusive(x_13); -if (x_16 == 0) +lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; +x_1 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1; +x_2 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2; +x_3 = lean_unsigned_to_nat(16u); +x_4 = lean_unsigned_to_nat(14u); +x_5 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3; +x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); +return x_6; +} +} +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +_start: { -lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; -x_17 = lean_ctor_get(x_13, 0); -x_18 = lean_ctor_get(x_13, 1); -x_19 = lean_ctor_get(x_13, 2); -x_20 = lean_ctor_get(x_13, 3); -x_21 = lean_ctor_get(x_13, 4); -x_22 = lean_ctor_get(x_13, 5); -x_23 = lean_ctor_get(x_13, 6); -x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1; -x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2; -x_26 = 0; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -x_27 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_22, x_24, x_25, x_26, x_26, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_27) == 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_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); -lean_inc(x_29); -lean_dec(x_27); -lean_ctor_set(x_13, 5, x_28); -x_30 = 1; -x_31 = lean_usize_add(x_2, x_30); -x_32 = lean_array_uset(x_15, x_2, x_13); -x_2 = x_31; -x_3 = x_32; -x_10 = x_29; +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; +x_18 = lean_unsigned_to_nat(1u); +x_19 = lean_nat_sub(x_5, x_18); +lean_dec(x_5); +x_20 = lean_array_fget(x_4, x_6); +x_21 = lean_array_get_size(x_2); +x_22 = lean_nat_dec_lt(x_6, 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; +x_23 = lean_ctor_get(x_20, 6); +lean_inc(x_23); +x_24 = lean_ctor_get(x_20, 3); +lean_inc(x_24); +x_25 = lean_ctor_get(x_20, 4); +lean_inc(x_25); +lean_dec(x_20); +x_26 = lean_ctor_get(x_23, 4); +lean_inc(x_26); +x_27 = lean_ctor_get(x_23, 2); +lean_inc(x_27); +lean_dec(x_23); +x_28 = l_Lean_Elab_withCommonTelescope___rarg___closed__1; +x_29 = l___private_Init_GetElem_0__outOfBounds___rarg(x_28); +x_30 = lean_array_get_size(x_29); +lean_dec(x_29); +x_31 = lean_nat_add(x_1, x_30); +lean_dec(x_30); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4; +x_33 = l_panic___at_Lean_Elab_wfRecursion___spec__9(x_32); +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_34 = l_Lean_Elab_WF_TerminationArgument_elab(x_24, x_25, x_31, x_26, x_33, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +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, 0); +lean_inc(x_35); +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = lean_nat_add(x_6, x_18); +lean_dec(x_6); +x_38 = lean_array_push(x_8, x_35); +x_5 = x_19; +x_6 = x_37; +x_7 = lean_box(0); +x_8 = x_38; +x_15 = x_36; goto _start; } else { -uint8_t x_34; -lean_free_object(x_13); -lean_dec(x_23); -lean_dec(x_21); -lean_dec(x_20); +uint8_t x_40; lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_17); -lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); lean_dec(x_6); -x_34 = !lean_is_exclusive(x_27); -if (x_34 == 0) +x_40 = !lean_is_exclusive(x_34); +if (x_40 == 0) { -return x_27; +return x_34; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_27, 0); -x_36 = lean_ctor_get(x_27, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_27); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_34, 0); +x_42 = lean_ctor_get(x_34, 1); +lean_inc(x_42); +lean_inc(x_41); +lean_dec(x_34); +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_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; uint8_t x_48; lean_object* x_49; -x_38 = lean_ctor_get(x_13, 0); -x_39 = lean_ctor_get_uint8(x_13, sizeof(void*)*7); -x_40 = lean_ctor_get(x_13, 1); -x_41 = lean_ctor_get(x_13, 2); -x_42 = lean_ctor_get(x_13, 3); -x_43 = lean_ctor_get(x_13, 4); -x_44 = lean_ctor_get(x_13, 5); -x_45 = lean_ctor_get(x_13, 6); -lean_inc(x_45); +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_27, 0); lean_inc(x_44); -lean_inc(x_43); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_38); -lean_dec(x_13); -x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1; -x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2; -x_48 = 0; +lean_dec(x_27); +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_49 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_44, x_46, x_47, x_48, x_48, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_49) == 0) +x_45 = l_Lean_Elab_WF_TerminationArgument_elab(x_24, x_25, x_31, x_26, x_44, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_45) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; -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_alloc_ctor(0, 7, 1); -lean_ctor_set(x_52, 0, x_38); -lean_ctor_set(x_52, 1, x_40); -lean_ctor_set(x_52, 2, x_41); -lean_ctor_set(x_52, 3, x_42); -lean_ctor_set(x_52, 4, x_43); -lean_ctor_set(x_52, 5, x_50); -lean_ctor_set(x_52, 6, x_45); -lean_ctor_set_uint8(x_52, sizeof(void*)*7, x_39); -x_53 = 1; -x_54 = lean_usize_add(x_2, x_53); -x_55 = lean_array_uset(x_15, x_2, x_52); -x_2 = x_54; -x_3 = x_55; -x_10 = x_51; +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; +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_nat_add(x_6, x_18); +lean_dec(x_6); +x_49 = lean_array_push(x_8, x_46); +x_5 = x_19; +x_6 = x_48; +x_7 = lean_box(0); +x_8 = x_49; +x_15 = x_47; goto _start; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; -lean_dec(x_45); -lean_dec(x_43); -lean_dec(x_42); -lean_dec(x_41); -lean_dec(x_40); -lean_dec(x_38); -lean_dec(x_15); +uint8_t x_51; +lean_dec(x_19); +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_57 = lean_ctor_get(x_49, 0); -lean_inc(x_57); -x_58 = lean_ctor_get(x_49, 1); -lean_inc(x_58); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_59 = x_49; -} else { - lean_dec_ref(x_49); - x_59 = lean_box(0); -} -if (lean_is_scalar(x_59)) { - x_60 = lean_alloc_ctor(1, 2, 0); -} else { - x_60 = x_59; -} -lean_ctor_set(x_60, 0, x_57); -lean_ctor_set(x_60, 1, x_58); -return x_60; -} -} -} -} +x_51 = !lean_is_exclusive(x_45); +if (x_51 == 0) +{ +return x_45; } -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__1___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__1___boxed), 4, 0); -return x_1; +lean_object* x_52; lean_object* x_53; lean_object* x_54; +x_52 = lean_ctor_get(x_45, 0); +x_53 = lean_ctor_get(x_45, 1); +lean_inc(x_53); +lean_inc(x_52); +lean_dec(x_45); +x_54 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_54, 0, x_52); +lean_ctor_set(x_54, 1, x_53); +return x_54; } } -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__1___closed__2() { -_start: -{ -lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__2___boxed), 4, 0); -return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_18 = lean_st_ref_get(x_16, 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_ctor_get(x_19, 0); -lean_inc(x_21); -lean_dec(x_19); -lean_inc(x_16); -lean_inc(x_15); -x_22 = l_Lean_Elab_addAsAxiom(x_1, x_13, x_14, x_15, x_16, x_20); -if (lean_obj_tag(x_22) == 0) +else { -lean_object* x_23; lean_object* x_24; lean_object* x_25; -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(x_2, x_3, x_4); -lean_inc(x_16); -lean_inc(x_15); +lean_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 = lean_ctor_get(x_20, 6); +lean_inc(x_55); +x_56 = lean_ctor_get(x_20, 3); +lean_inc(x_56); +x_57 = lean_ctor_get(x_20, 4); +lean_inc(x_57); +lean_dec(x_20); +x_58 = lean_ctor_get(x_55, 4); +lean_inc(x_58); +x_59 = lean_ctor_get(x_55, 2); +lean_inc(x_59); +lean_dec(x_55); +x_60 = lean_array_fget(x_2, x_6); +x_61 = lean_array_get_size(x_60); +lean_dec(x_60); +x_62 = lean_nat_add(x_1, x_61); +lean_dec(x_61); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4; +x_64 = l_panic___at_Lean_Elab_wfRecursion___spec__9(x_63); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); -lean_inc(x_1); -x_25 = l_Lean_Elab_WF_mkFix(x_1, x_5, x_6, x_7, x_24, x_11, x_12, x_13, x_14, x_15, x_16, x_23); -if (lean_obj_tag(x_25) == 0) +lean_inc(x_10); +lean_inc(x_9); +x_65 = l_Lean_Elab_WF_TerminationArgument_elab(x_56, x_57, x_62, x_58, x_64, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_65) == 0) { -lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_26 = lean_ctor_get(x_25, 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_wfRecursion___lambda__1___closed__1; -x_29 = l_Lean_Elab_wfRecursion___lambda__1___closed__2; -lean_inc(x_16); -lean_inc(x_15); -x_30 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_26, x_28, x_29, x_15, x_16, x_27); -if (lean_obj_tag(x_30) == 0) +lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_66 = lean_ctor_get(x_65, 0); +lean_inc(x_66); +x_67 = lean_ctor_get(x_65, 1); +lean_inc(x_67); +lean_dec(x_65); +x_68 = lean_nat_add(x_6, x_18); +lean_dec(x_6); +x_69 = lean_array_push(x_8, x_66); +x_5 = x_19; +x_6 = x_68; +x_7 = lean_box(0); +x_8 = x_69; +x_15 = x_67; +goto _start; +} +else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_31 = lean_ctor_get(x_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_st_ref_get(x_16, x_32); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_ctor_get(x_34, 0); -lean_inc(x_36); -lean_dec(x_34); -x_37 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_35); +uint8_t x_71; +lean_dec(x_19); lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -x_38 = lean_ctor_get(x_37, 1); -lean_inc(x_38); -lean_dec(x_37); -x_39 = l_Lean_Meta_unfoldDeclsFrom(x_36, x_31, x_15, x_16, x_38); -if (lean_obj_tag(x_39) == 0) -{ -uint8_t x_40; -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_6); +x_71 = !lean_is_exclusive(x_65); +if (x_71 == 0) { -lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_41 = lean_ctor_get(x_39, 0); -x_42 = lean_ctor_get(x_1, 0); -lean_inc(x_42); -x_43 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -x_44 = lean_ctor_get(x_1, 1); -lean_inc(x_44); -x_45 = lean_ctor_get(x_1, 2); -lean_inc(x_45); -x_46 = lean_ctor_get(x_1, 6); -lean_inc(x_46); -lean_dec(x_1); -x_47 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_47, 0, x_42); -lean_ctor_set(x_47, 1, x_44); -lean_ctor_set(x_47, 2, x_45); -lean_ctor_set(x_47, 3, x_8); -lean_ctor_set(x_47, 4, x_9); -lean_ctor_set(x_47, 5, x_41); -lean_ctor_set(x_47, 6, x_46); -lean_ctor_set_uint8(x_47, sizeof(void*)*7, x_43); -lean_ctor_set(x_39, 0, x_47); -return x_39; +return x_65; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_48 = lean_ctor_get(x_39, 0); -x_49 = lean_ctor_get(x_39, 1); -lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_39); -x_50 = lean_ctor_get(x_1, 0); -lean_inc(x_50); -x_51 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); -x_52 = lean_ctor_get(x_1, 1); -lean_inc(x_52); -x_53 = lean_ctor_get(x_1, 2); -lean_inc(x_53); -x_54 = lean_ctor_get(x_1, 6); -lean_inc(x_54); -lean_dec(x_1); -x_55 = lean_alloc_ctor(0, 7, 1); -lean_ctor_set(x_55, 0, x_50); -lean_ctor_set(x_55, 1, x_52); -lean_ctor_set(x_55, 2, x_53); -lean_ctor_set(x_55, 3, x_8); -lean_ctor_set(x_55, 4, x_9); -lean_ctor_set(x_55, 5, x_48); -lean_ctor_set(x_55, 6, x_54); -lean_ctor_set_uint8(x_55, sizeof(void*)*7, x_51); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_49); -return x_56; +lean_object* x_72; lean_object* x_73; lean_object* x_74; +x_72 = lean_ctor_get(x_65, 0); +x_73 = lean_ctor_get(x_65, 1); +lean_inc(x_73); +lean_inc(x_72); +lean_dec(x_65); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_72); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} } } else { -uint8_t x_57; +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_59, 0); +lean_inc(x_75); +lean_dec(x_59); +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_76 = l_Lean_Elab_WF_TerminationArgument_elab(x_56, x_57, x_62, x_58, x_75, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +if (lean_obj_tag(x_76) == 0) +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; +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_nat_add(x_6, x_18); +lean_dec(x_6); +x_80 = lean_array_push(x_8, x_77); +x_5 = x_19; +x_6 = x_79; +x_7 = lean_box(0); +x_8 = x_80; +x_15 = x_78; +goto _start; +} +else +{ +uint8_t x_82; +lean_dec(x_19); +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_1); -x_57 = !lean_is_exclusive(x_39); -if (x_57 == 0) +lean_dec(x_6); +x_82 = !lean_is_exclusive(x_76); +if (x_82 == 0) { -return x_39; +return x_76; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_39, 0); -x_59 = lean_ctor_get(x_39, 1); -lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_39); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_object* x_83; lean_object* x_84; lean_object* x_85; +x_83 = lean_ctor_get(x_76, 0); +x_84 = lean_ctor_get(x_76, 1); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_76); +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 { -lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_1); -x_61 = lean_ctor_get(x_30, 0); -lean_inc(x_61); -x_62 = lean_ctor_get(x_30, 1); -lean_inc(x_62); -lean_dec(x_30); -x_63 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_62); -lean_dec(x_16); -lean_dec(x_15); +lean_object* x_86; lean_dec(x_14); lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); -x_64 = !lean_is_exclusive(x_63); -if (x_64 == 0) +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_6); +lean_dec(x_5); +x_86 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_86, 0, x_8); +lean_ctor_set(x_86, 1, x_15); +return x_86; +} +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_object* x_10, lean_object* x_11) { +_start: { -lean_object* x_65; +uint8_t x_12; +x_12 = lean_usize_dec_lt(x_3, x_2); +if (x_12 == 0) +{ +lean_object* x_13; +lean_dec(x_10); +lean_dec(x_9); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_4); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; +lean_dec(x_4); +x_14 = lean_array_uget(x_1, x_3); +lean_inc(x_10); +lean_inc(x_9); +x_15 = l_Lean_Elab_addAsAxiom(x_14, x_7, x_8, x_9, x_10, x_11); +lean_dec(x_14); +if (lean_obj_tag(x_15) == 0) +{ +lean_object* x_16; size_t x_17; size_t x_18; lean_object* x_19; +x_16 = lean_ctor_get(x_15, 1); +lean_inc(x_16); +lean_dec(x_15); +x_17 = 1; +x_18 = lean_usize_add(x_3, x_17); +x_19 = lean_box(0); +x_3 = x_18; +x_4 = x_19; +x_11 = x_16; +goto _start; +} +else +{ +uint8_t x_21; +lean_dec(x_10); +lean_dec(x_9); +x_21 = !lean_is_exclusive(x_15); +if (x_21 == 0) +{ +return x_15; +} +else +{ +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_15, 0); +x_23 = lean_ctor_get(x_15, 1); +lean_inc(x_23); +lean_inc(x_22); +lean_dec(x_15); +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_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__12(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; +x_12 = lean_usize_dec_lt(x_3, x_2); +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_1); +x_13 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_13, 0, x_4); +lean_ctor_set(x_13, 1, x_11); +return x_13; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_14 = lean_array_uget(x_4, x_3); +x_15 = lean_unsigned_to_nat(0u); +x_16 = lean_array_uset(x_4, x_3, x_15); +lean_inc(x_10); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_1); +x_17 = l_Lean_Elab_varyingVarNames(x_1, x_14, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; size_t x_20; size_t x_21; lean_object* x_22; +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 = 1; +x_21 = lean_usize_add(x_3, x_20); +x_22 = lean_array_uset(x_16, x_3, x_18); +x_3 = x_21; +x_4 = x_22; +x_11 = x_19; +goto _start; +} +else +{ +uint8_t x_24; +lean_dec(x_16); +lean_dec(x_10); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_1); +x_24 = !lean_is_exclusive(x_17); +if (x_24 == 0) +{ +return x_17; +} +else +{ +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); +lean_inc(x_25); +lean_dec(x_17); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; +} +} +} +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_iteToDIte___lambda__2___boxed), 6, 0); +return x_1; +} +} +static lean_object* _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_iteToDIte___lambda__1___boxed), 6, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13(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; +x_11 = lean_usize_dec_lt(x_2, x_1); +if (x_11 == 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_3); +lean_ctor_set(x_12, 1, x_10); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_13 = lean_array_uget(x_3, x_2); +x_14 = lean_unsigned_to_nat(0u); +x_15 = lean_array_uset(x_3, x_2, x_14); +x_16 = !lean_is_exclusive(x_13); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; lean_object* x_27; +x_17 = lean_ctor_get(x_13, 0); +x_18 = lean_ctor_get(x_13, 1); +x_19 = lean_ctor_get(x_13, 2); +x_20 = lean_ctor_get(x_13, 3); +x_21 = lean_ctor_get(x_13, 4); +x_22 = lean_ctor_get(x_13, 5); +x_23 = lean_ctor_get(x_13, 6); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1; +x_25 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2; +x_26 = 0; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_27 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_22, x_24, x_25, x_26, x_26, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_27) == 0) +{ +lean_object* x_28; lean_object* x_29; size_t x_30; size_t x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_27, 0); +lean_inc(x_28); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +lean_ctor_set(x_13, 5, x_28); +x_30 = 1; +x_31 = lean_usize_add(x_2, x_30); +x_32 = lean_array_uset(x_15, x_2, x_13); +x_2 = x_31; +x_3 = x_32; +x_10 = x_29; +goto _start; +} +else +{ +uint8_t x_34; +lean_free_object(x_13); +lean_dec(x_23); +lean_dec(x_21); +lean_dec(x_20); +lean_dec(x_19); +lean_dec(x_18); +lean_dec(x_17); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_34 = !lean_is_exclusive(x_27); +if (x_34 == 0) +{ +return x_27; +} +else +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; +x_35 = lean_ctor_get(x_27, 0); +x_36 = lean_ctor_get(x_27, 1); +lean_inc(x_36); +lean_inc(x_35); +lean_dec(x_27); +x_37 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_37, 0, x_35); +lean_ctor_set(x_37, 1, x_36); +return x_37; +} +} +} +else +{ +lean_object* x_38; 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; uint8_t x_48; lean_object* x_49; +x_38 = lean_ctor_get(x_13, 0); +x_39 = lean_ctor_get_uint8(x_13, sizeof(void*)*7); +x_40 = lean_ctor_get(x_13, 1); +x_41 = lean_ctor_get(x_13, 2); +x_42 = lean_ctor_get(x_13, 3); +x_43 = lean_ctor_get(x_13, 4); +x_44 = lean_ctor_get(x_13, 5); +x_45 = lean_ctor_get(x_13, 6); +lean_inc(x_45); +lean_inc(x_44); +lean_inc(x_43); +lean_inc(x_42); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_38); +lean_dec(x_13); +x_46 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1; +x_47 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2; +x_48 = 0; +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +x_49 = l_Lean_Meta_transform___at_Lean_Meta_zetaReduce___spec__1(x_44, x_46, x_47, x_48, x_48, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_49) == 0) +{ +lean_object* x_50; lean_object* x_51; lean_object* x_52; size_t x_53; size_t x_54; lean_object* x_55; +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_alloc_ctor(0, 7, 1); +lean_ctor_set(x_52, 0, x_38); +lean_ctor_set(x_52, 1, x_40); +lean_ctor_set(x_52, 2, x_41); +lean_ctor_set(x_52, 3, x_42); +lean_ctor_set(x_52, 4, x_43); +lean_ctor_set(x_52, 5, x_50); +lean_ctor_set(x_52, 6, x_45); +lean_ctor_set_uint8(x_52, sizeof(void*)*7, x_39); +x_53 = 1; +x_54 = lean_usize_add(x_2, x_53); +x_55 = lean_array_uset(x_15, x_2, x_52); +x_2 = x_54; +x_3 = x_55; +x_10 = x_51; +goto _start; +} +else +{ +lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +lean_dec(x_45); +lean_dec(x_43); +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_40); +lean_dec(x_38); +lean_dec(x_15); +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +x_57 = lean_ctor_get(x_49, 0); +lean_inc(x_57); +x_58 = lean_ctor_get(x_49, 1); +lean_inc(x_58); +if (lean_is_exclusive(x_49)) { + lean_ctor_release(x_49, 0); + lean_ctor_release(x_49, 1); + x_59 = x_49; +} else { + lean_dec_ref(x_49); + x_59 = lean_box(0); +} +if (lean_is_scalar(x_59)) { + x_60 = lean_alloc_ctor(1, 2, 0); +} else { + x_60 = x_59; +} +lean_ctor_set(x_60, 0, x_57); +lean_ctor_set(x_60, 1, x_58); +return x_60; +} +} +} +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__1___boxed), 4, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_eraseRecAppSyntaxExpr___lambda__2___boxed), 4, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_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_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_18 = lean_st_ref_get(x_16, 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_ctor_get(x_19, 0); +lean_inc(x_21); +lean_dec(x_19); +lean_inc(x_16); +lean_inc(x_15); +x_22 = l_Lean_Elab_addAsAxiom(x_1, x_13, x_14, x_15, x_16, x_20); +if (lean_obj_tag(x_22) == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); +x_24 = l_Array_mapMUnsafe_map___at_Lean_Elab_WF_guessLex___spec__4(x_2, x_3, x_4); +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_1); +x_25 = l_Lean_Elab_WF_mkFix(x_1, x_5, x_6, x_7, x_24, x_11, x_12, x_13, x_14, x_15, x_16, x_23); +if (lean_obj_tag(x_25) == 0) +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +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_wfRecursion___lambda__1___closed__1; +x_29 = l_Lean_Elab_wfRecursion___lambda__1___closed__2; +lean_inc(x_16); +lean_inc(x_15); +x_30 = l_Lean_Core_transform___at_Lean_Core_betaReduce___spec__1(x_26, x_28, x_29, x_15, x_16, 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; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = lean_st_ref_get(x_16, x_32); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = lean_ctor_get(x_34, 0); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_35); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_38 = lean_ctor_get(x_37, 1); +lean_inc(x_38); +lean_dec(x_37); +x_39 = l_Lean_Meta_unfoldDeclsFrom(x_36, x_31, x_15, x_16, x_38); +if (lean_obj_tag(x_39) == 0) +{ +uint8_t x_40; +x_40 = !lean_is_exclusive(x_39); +if (x_40 == 0) +{ +lean_object* x_41; lean_object* x_42; uint8_t x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_41 = lean_ctor_get(x_39, 0); +x_42 = lean_ctor_get(x_1, 0); +lean_inc(x_42); +x_43 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +x_44 = lean_ctor_get(x_1, 1); +lean_inc(x_44); +x_45 = lean_ctor_get(x_1, 2); +lean_inc(x_45); +x_46 = lean_ctor_get(x_1, 6); +lean_inc(x_46); +lean_dec(x_1); +x_47 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_47, 0, x_42); +lean_ctor_set(x_47, 1, x_44); +lean_ctor_set(x_47, 2, x_45); +lean_ctor_set(x_47, 3, x_8); +lean_ctor_set(x_47, 4, x_9); +lean_ctor_set(x_47, 5, x_41); +lean_ctor_set(x_47, 6, x_46); +lean_ctor_set_uint8(x_47, sizeof(void*)*7, x_43); +lean_ctor_set(x_39, 0, x_47); +return x_39; +} +else +{ +lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_48 = lean_ctor_get(x_39, 0); +x_49 = lean_ctor_get(x_39, 1); +lean_inc(x_49); +lean_inc(x_48); +lean_dec(x_39); +x_50 = lean_ctor_get(x_1, 0); +lean_inc(x_50); +x_51 = lean_ctor_get_uint8(x_1, sizeof(void*)*7); +x_52 = lean_ctor_get(x_1, 1); +lean_inc(x_52); +x_53 = lean_ctor_get(x_1, 2); +lean_inc(x_53); +x_54 = lean_ctor_get(x_1, 6); +lean_inc(x_54); +lean_dec(x_1); +x_55 = lean_alloc_ctor(0, 7, 1); +lean_ctor_set(x_55, 0, x_50); +lean_ctor_set(x_55, 1, x_52); +lean_ctor_set(x_55, 2, x_53); +lean_ctor_set(x_55, 3, x_8); +lean_ctor_set(x_55, 4, x_9); +lean_ctor_set(x_55, 5, x_48); +lean_ctor_set(x_55, 6, x_54); +lean_ctor_set_uint8(x_55, sizeof(void*)*7, x_51); +x_56 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_56, 0, x_55); +lean_ctor_set(x_56, 1, x_49); +return x_56; +} +} +else +{ +uint8_t x_57; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_57 = !lean_is_exclusive(x_39); +if (x_57 == 0) +{ +return x_39; +} +else +{ +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_39, 0); +x_59 = lean_ctor_get(x_39, 1); +lean_inc(x_59); +lean_inc(x_58); +lean_dec(x_39); +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_58); +lean_ctor_set(x_60, 1, x_59); +return x_60; +} +} +} +else +{ +lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_61 = lean_ctor_get(x_30, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_30, 1); +lean_inc(x_62); +lean_dec(x_30); +x_63 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_62); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_64 = !lean_is_exclusive(x_63); +if (x_64 == 0) +{ +lean_object* x_65; x_65 = lean_ctor_get(x_63, 0); lean_dec(x_65); lean_ctor_set_tag(x_63, 1); @@ -7592,603 +8200,981 @@ return x_63; } else { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_67, 0, x_61); -lean_ctor_set(x_67, 1, x_66); -return x_67; +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); +x_67 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_67, 0, x_61); +lean_ctor_set(x_67, 1, x_66); +return x_67; +} +} +} +else +{ +lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +lean_dec(x_9); +lean_dec(x_8); +lean_dec(x_1); +x_68 = lean_ctor_get(x_25, 0); +lean_inc(x_68); +x_69 = lean_ctor_get(x_25, 1); +lean_inc(x_69); +lean_dec(x_25); +x_70 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_69); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_71 = !lean_is_exclusive(x_70); +if (x_71 == 0) +{ +lean_object* x_72; +x_72 = lean_ctor_get(x_70, 0); +lean_dec(x_72); +lean_ctor_set_tag(x_70, 1); +lean_ctor_set(x_70, 0, x_68); +return x_70; +} +else +{ +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_70, 1); +lean_inc(x_73); +lean_dec(x_70); +x_74 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_74, 0, x_68); +lean_ctor_set(x_74, 1, x_73); +return x_74; +} +} +} +else +{ +lean_object* x_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +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_75 = lean_ctor_get(x_22, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_22, 1); +lean_inc(x_76); +lean_dec(x_22); +x_77 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_76); +lean_dec(x_16); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_78 = !lean_is_exclusive(x_77); +if (x_78 == 0) +{ +lean_object* x_79; +x_79 = lean_ctor_get(x_77, 0); +lean_dec(x_79); +lean_ctor_set_tag(x_77, 1); +lean_ctor_set(x_77, 0, x_75); +return x_77; +} +else +{ +lean_object* x_80; lean_object* x_81; +x_80 = lean_ctor_get(x_77, 1); +lean_inc(x_80); +lean_dec(x_77); +x_81 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_81, 0, x_75); +lean_ctor_set(x_81, 1, x_80); +return x_81; +} +} +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__2___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("wfRel: ", 7); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__2___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_wfRecursion___lambda__2___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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, lean_object* x_16) { +_start: +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_17 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__4; +x_18 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_17, x_10, x_11, x_12, x_13, x_14, x_15, x_16); +x_19 = lean_ctor_get(x_18, 0); +lean_inc(x_19); +x_20 = lean_unbox(x_19); +lean_dec(x_19); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +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_wfRecursion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_7, x_8, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_21); +return x_23; +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_24 = lean_ctor_get(x_18, 1); +lean_inc(x_24); +lean_dec(x_18); +lean_inc(x_9); +x_25 = l_Lean_MessageData_ofExpr(x_9); +x_26 = l_Lean_Elab_wfRecursion___lambda__2___closed__2; +x_27 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_27, 0, x_26); +lean_ctor_set(x_27, 1, x_25); +x_28 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__6; +x_29 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +x_30 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_17, x_29, x_10, x_11, x_12, x_13, x_14, x_15, x_24); +x_31 = lean_ctor_get(x_30, 0); +lean_inc(x_31); +x_32 = lean_ctor_get(x_30, 1); +lean_inc(x_32); +lean_dec(x_30); +x_33 = l_Lean_Elab_wfRecursion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_7, x_8, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_32); +lean_dec(x_31); +return x_33; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { +_start: +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = l_Lean_Expr_bindingDomain_x21(x_1); +x_19 = lean_ctor_get(x_2, 3); +lean_inc(x_19); +x_20 = lean_box_usize(x_3); +x_21 = lean_box_usize(x_4); +lean_inc(x_19); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +x_22 = lean_alloc_closure((void*)(l_Lean_Elab_wfRecursion___lambda__2___boxed), 16, 8); +lean_closure_set(x_22, 0, x_2); +lean_closure_set(x_22, 1, x_20); +lean_closure_set(x_22, 2, x_21); +lean_closure_set(x_22, 3, x_5); +lean_closure_set(x_22, 4, x_6); +lean_closure_set(x_22, 5, x_7); +lean_closure_set(x_22, 6, x_19); +lean_closure_set(x_22, 7, x_8); +x_23 = l_Lean_Elab_WF_elabWFRel___rarg(x_5, x_19, x_6, x_7, x_18, x_9, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_17); +return x_23; +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__4___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("wfRecursion: expected unary function type: ", 43); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__4___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_wfRecursion___lambda__4___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16) { +_start: +{ +lean_object* x_17; +lean_inc(x_15); +lean_inc(x_14); +lean_inc(x_13); +lean_inc(x_12); +x_17 = l_Lean_Meta_whnfForall(x_9, x_12, x_13, x_14, x_15, x_16); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); +lean_inc(x_19); +lean_dec(x_17); +x_20 = l_Lean_Expr_isForall(x_18); +if (x_20 == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_21 = l_Lean_MessageData_ofExpr(x_18); +x_22 = l_Lean_Elab_wfRecursion___lambda__4___closed__2; +x_23 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_21); +x_24 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__6; +x_25 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +x_26 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_13); +lean_dec(x_12); +lean_dec(x_11); +x_27 = !lean_is_exclusive(x_26); +if (x_27 == 0) +{ +return x_26; +} +else +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_28 = lean_ctor_get(x_26, 0); +x_29 = lean_ctor_get(x_26, 1); +lean_inc(x_29); +lean_inc(x_28); +lean_dec(x_26); +x_30 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_30, 0, x_28); +lean_ctor_set(x_30, 1, x_29); +return x_30; +} +} +else +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_box(0); +x_32 = l_Lean_Elab_wfRecursion___lambda__3(x_18, x_1, x_2, x_3, x_4, x_8, x_5, x_6, x_7, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +lean_dec(x_18); +return x_32; +} +} +else +{ +uint8_t x_33; +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_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +x_33 = !lean_is_exclusive(x_17); +if (x_33 == 0) +{ +return x_17; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_17, 0); +x_35 = lean_ctor_get(x_17, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_17); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; +} } } } +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__5___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_maxRecDepth; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__5(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) { +_start: +{ +uint8_t x_15; +x_15 = !lean_is_exclusive(x_12); +if (x_15 == 0) +{ +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_16 = lean_ctor_get(x_12, 4); +lean_dec(x_16); +x_17 = lean_ctor_get(x_12, 2); +lean_dec(x_17); +x_18 = l_Lean_Elab_wfRecursion___lambda__5___closed__1; +x_19 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_18); +lean_ctor_set(x_12, 4, x_19); +lean_ctor_set(x_12, 2, x_1); +lean_ctor_set_uint8(x_12, sizeof(void*)*11 + 1, x_2); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_20 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef(x_3, x_4, x_5, x_6, x_12, x_13, x_14); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; uint8_t x_22; +x_21 = lean_ctor_get(x_20, 0); +lean_inc(x_21); +x_22 = lean_unbox(x_21); +lean_dec(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_23 = lean_ctor_get(x_20, 1); +lean_inc(x_23); +lean_dec(x_20); +x_24 = lean_ctor_get(x_7, 3); +lean_inc(x_24); +x_25 = lean_box(0); +x_26 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_26, 0, x_24); +lean_ctor_set(x_26, 1, x_25); +x_27 = 0; +x_28 = lean_box(x_27); +lean_inc(x_7); +x_29 = lean_alloc_closure((void*)(l_Lean_Elab_addNonRec___boxed), 10, 3); +lean_closure_set(x_29, 0, x_7); +lean_closure_set(x_29, 1, x_28); +lean_closure_set(x_29, 2, x_26); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_9); +lean_inc(x_8); +x_30 = l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6(x_27, x_29, x_8, x_9, x_5, x_6, x_12, x_13, x_23); +if (lean_obj_tag(x_30) == 0) +{ +lean_object* x_31; lean_object* x_32; +x_31 = lean_ctor_get(x_30, 1); +lean_inc(x_31); +lean_dec(x_30); +x_32 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs(x_4, x_10, x_3, x_7, x_8, x_9, x_5, x_6, x_12, x_13, x_31); +return x_32; +} else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; uint8_t x_71; +uint8_t x_33; +lean_dec(x_12); +lean_dec(x_13); +lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_1); -x_68 = lean_ctor_get(x_25, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_25, 1); -lean_inc(x_69); -lean_dec(x_25); -x_70 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_69); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -x_71 = !lean_is_exclusive(x_70); -if (x_71 == 0) +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_30); +if (x_33 == 0) { -lean_object* x_72; -x_72 = lean_ctor_get(x_70, 0); -lean_dec(x_72); -lean_ctor_set_tag(x_70, 1); -lean_ctor_set(x_70, 0, x_68); -return x_70; +return x_30; } else { -lean_object* x_73; lean_object* x_74; -x_73 = lean_ctor_get(x_70, 1); -lean_inc(x_73); -lean_dec(x_70); -x_74 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_74, 0, x_68); -lean_ctor_set(x_74, 1, x_73); -return x_74; +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 0); +x_35 = lean_ctor_get(x_30, 1); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_30); +x_36 = lean_alloc_ctor(1, 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_75; lean_object* x_76; lean_object* x_77; uint8_t x_78; +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_t x_41; lean_object* x_42; +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_3); +x_37 = lean_ctor_get(x_20, 1); +lean_inc(x_37); +lean_dec(x_20); +x_38 = lean_ctor_get(x_7, 3); +lean_inc(x_38); +x_39 = lean_box(0); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +x_41 = 0; +x_42 = l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux(x_7, x_41, x_40, x_41, x_8, x_9, x_5, x_6, x_12, x_13, x_37); +return x_42; +} +} +else +{ +uint8_t x_43; +lean_dec(x_12); +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_1); -x_75 = lean_ctor_get(x_22, 0); -lean_inc(x_75); -x_76 = lean_ctor_get(x_22, 1); -lean_inc(x_76); -lean_dec(x_22); -x_77 = l_Lean_setEnv___at_Lean_Elab_Term_evalTerm___spec__1(x_21, x_11, x_12, x_13, x_14, x_15, x_16, x_76); -lean_dec(x_16); -lean_dec(x_15); -lean_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -x_78 = !lean_is_exclusive(x_77); -if (x_78 == 0) +lean_dec(x_3); +x_43 = !lean_is_exclusive(x_20); +if (x_43 == 0) { -lean_object* x_79; -x_79 = lean_ctor_get(x_77, 0); -lean_dec(x_79); -lean_ctor_set_tag(x_77, 1); -lean_ctor_set(x_77, 0, x_75); -return x_77; +return x_20; } else { -lean_object* x_80; lean_object* x_81; -x_80 = lean_ctor_get(x_77, 1); -lean_inc(x_80); -lean_dec(x_77); -x_81 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_81, 0, x_75); -lean_ctor_set(x_81, 1, x_80); -return x_81; -} +lean_object* x_44; lean_object* x_45; lean_object* x_46; +x_44 = lean_ctor_get(x_20, 0); +x_45 = lean_ctor_get(x_20, 1); +lean_inc(x_45); +lean_inc(x_44); +lean_dec(x_20); +x_46 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_45); +return x_46; } } } -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__2___closed__1() { -_start: +else { -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("wfRel: ", 7); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__2___closed__2() { -_start: +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_47 = lean_ctor_get(x_12, 0); +x_48 = lean_ctor_get(x_12, 1); +x_49 = lean_ctor_get(x_12, 3); +x_50 = lean_ctor_get(x_12, 5); +x_51 = lean_ctor_get(x_12, 6); +x_52 = lean_ctor_get(x_12, 7); +x_53 = lean_ctor_get(x_12, 8); +x_54 = lean_ctor_get(x_12, 9); +x_55 = lean_ctor_get(x_12, 10); +x_56 = lean_ctor_get_uint8(x_12, sizeof(void*)*11); +lean_inc(x_55); +lean_inc(x_54); +lean_inc(x_53); +lean_inc(x_52); +lean_inc(x_51); +lean_inc(x_50); +lean_inc(x_49); +lean_inc(x_48); +lean_inc(x_47); +lean_dec(x_12); +x_57 = l_Lean_Elab_wfRecursion___lambda__5___closed__1; +x_58 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_57); +x_59 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_59, 0, x_47); +lean_ctor_set(x_59, 1, x_48); +lean_ctor_set(x_59, 2, x_1); +lean_ctor_set(x_59, 3, x_49); +lean_ctor_set(x_59, 4, x_58); +lean_ctor_set(x_59, 5, x_50); +lean_ctor_set(x_59, 6, x_51); +lean_ctor_set(x_59, 7, x_52); +lean_ctor_set(x_59, 8, x_53); +lean_ctor_set(x_59, 9, x_54); +lean_ctor_set(x_59, 10, x_55); +lean_ctor_set_uint8(x_59, sizeof(void*)*11, x_56); +lean_ctor_set_uint8(x_59, sizeof(void*)*11 + 1, x_2); +lean_inc(x_13); +lean_inc(x_59); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_60 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef(x_3, x_4, x_5, x_6, x_59, x_13, x_14); +if (lean_obj_tag(x_60) == 0) { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_wfRecursion___lambda__2___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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, lean_object* x_16) { -_start: +lean_object* x_61; uint8_t x_62; +x_61 = lean_ctor_get(x_60, 0); +lean_inc(x_61); +x_62 = lean_unbox(x_61); +lean_dec(x_61); +if (x_62 == 0) { -lean_object* x_17; lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_17 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__4; -x_18 = l_Lean_isTracingEnabledFor___at_Lean_Elab_Term_traceAtCmdPos___spec__1(x_17, x_10, x_11, x_12, x_13, x_14, x_15, x_16); -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_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; +x_63 = lean_ctor_get(x_60, 1); +lean_inc(x_63); +lean_dec(x_60); +x_64 = lean_ctor_get(x_7, 3); +lean_inc(x_64); +x_65 = lean_box(0); +x_66 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_66, 0, x_64); +lean_ctor_set(x_66, 1, x_65); +x_67 = 0; +x_68 = lean_box(x_67); +lean_inc(x_7); +x_69 = lean_alloc_closure((void*)(l_Lean_Elab_addNonRec___boxed), 10, 3); +lean_closure_set(x_69, 0, x_7); +lean_closure_set(x_69, 1, x_68); +lean_closure_set(x_69, 2, x_66); +lean_inc(x_13); +lean_inc(x_59); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_9); +lean_inc(x_8); +x_70 = l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6(x_67, x_69, x_8, x_9, x_5, x_6, x_59, x_13, x_63); +if (lean_obj_tag(x_70) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -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_wfRecursion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_7, x_8, x_22, x_10, x_11, x_12, x_13, x_14, x_15, x_21); -return x_23; +lean_object* x_71; lean_object* x_72; +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +x_72 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs(x_4, x_10, x_3, x_7, x_8, x_9, x_5, x_6, x_59, x_13, x_71); +return x_72; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_24 = lean_ctor_get(x_18, 1); -lean_inc(x_24); -lean_dec(x_18); -lean_inc(x_9); -x_25 = l_Lean_MessageData_ofExpr(x_9); -x_26 = l_Lean_Elab_wfRecursion___lambda__2___closed__2; -x_27 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_27, 0, x_26); -lean_ctor_set(x_27, 1, x_25); -x_28 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__6; -x_29 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_29, 0, x_27); -lean_ctor_set(x_29, 1, x_28); -x_30 = l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(x_17, x_29, x_10, x_11, x_12, x_13, x_14, x_15, x_24); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Elab_wfRecursion___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_9, x_7, x_8, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_32); -lean_dec(x_31); -return x_33; +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; +lean_dec(x_59); +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); +x_73 = lean_ctor_get(x_70, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_70, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_70)) { + lean_ctor_release(x_70, 0); + lean_ctor_release(x_70, 1); + x_75 = x_70; +} else { + lean_dec_ref(x_70); + x_75 = lean_box(0); +} +if (lean_is_scalar(x_75)) { + x_76 = lean_alloc_ctor(1, 2, 0); +} else { + x_76 = x_75; } +lean_ctor_set(x_76, 0, x_73); +lean_ctor_set(x_76, 1, x_74); +return x_76; } } -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__3(lean_object* x_1, lean_object* x_2, size_t x_3, size_t x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15, lean_object* x_16, lean_object* x_17) { -_start: +else { -lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_18 = l_Lean_Expr_bindingDomain_x21(x_1); -x_19 = lean_ctor_get(x_2, 3); -lean_inc(x_19); -x_20 = lean_box_usize(x_3); -x_21 = lean_box_usize(x_4); -lean_inc(x_19); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_wfRecursion___lambda__2___boxed), 16, 8); -lean_closure_set(x_22, 0, x_2); -lean_closure_set(x_22, 1, x_20); -lean_closure_set(x_22, 2, x_21); -lean_closure_set(x_22, 3, x_5); -lean_closure_set(x_22, 4, x_6); -lean_closure_set(x_22, 5, x_7); -lean_closure_set(x_22, 6, x_19); -lean_closure_set(x_22, 7, x_8); -x_23 = l_Lean_Elab_WF_elabWFRel___rarg(x_5, x_19, x_6, x_7, x_18, x_9, x_22, x_11, x_12, x_13, x_14, x_15, x_16, x_17); -return x_23; +lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; uint8_t x_81; lean_object* x_82; +lean_dec(x_10); +lean_dec(x_4); +lean_dec(x_3); +x_77 = lean_ctor_get(x_60, 1); +lean_inc(x_77); +lean_dec(x_60); +x_78 = lean_ctor_get(x_7, 3); +lean_inc(x_78); +x_79 = lean_box(0); +x_80 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_80, 0, x_78); +lean_ctor_set(x_80, 1, x_79); +x_81 = 0; +x_82 = l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux(x_7, x_81, x_80, x_81, x_8, x_9, x_5, x_6, x_59, x_13, x_77); +return x_82; +} +} +else +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +lean_dec(x_59); +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); +x_83 = lean_ctor_get(x_60, 0); +lean_inc(x_83); +x_84 = lean_ctor_get(x_60, 1); +lean_inc(x_84); +if (lean_is_exclusive(x_60)) { + lean_ctor_release(x_60, 0); + lean_ctor_release(x_60, 1); + x_85 = x_60; +} else { + lean_dec_ref(x_60); + x_85 = lean_box(0); } +if (lean_is_scalar(x_85)) { + x_86 = lean_alloc_ctor(1, 2, 0); +} else { + x_86 = x_85; } -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__4___closed__1() { +lean_ctor_set(x_86, 0, x_83); +lean_ctor_set(x_86, 1, x_84); +return x_86; +} +} +} +} +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__6___closed__1() { _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("wfRecursion: expected unary function type: ", 43); +x_1 = l_Lean_allowUnsafeReducibility; return x_1; } } -static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Elab_wfRecursion___lambda__6___closed__2() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_wfRecursion___lambda__4___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; +lean_object* x_1; +x_1 = l_Lean_diagnostics; +return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__4(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, 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_wfRecursion___lambda__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { _start: { -lean_object* x_17; -lean_inc(x_15); -lean_inc(x_14); +lean_object* x_15; lean_inc(x_13); lean_inc(x_12); -x_17 = l_Lean_Meta_whnfForall(x_9, x_12, x_13, x_14, x_15, x_16); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; uint8_t x_20; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); -lean_inc(x_19); -lean_dec(x_17); -x_20 = l_Lean_Expr_isForall(x_18); -if (x_20 == 0) +x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_addAndCompileUnsafe___spec__1(x_1, x_2, x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(x_15) == 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; uint8_t x_27; -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_21 = l_Lean_MessageData_ofExpr(x_18); -x_22 = l_Lean_Elab_wfRecursion___lambda__4___closed__2; -x_23 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_23, 1, x_21); -x_24 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__6; -x_25 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_25, 0, x_23); -lean_ctor_set(x_25, 1, x_24); -x_26 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_25, x_10, x_11, x_12, x_13, x_14, x_15, x_19); +lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_54; lean_object* x_55; uint8_t x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; uint8_t x_64; uint8_t x_65; +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_dec(x_14); -lean_dec(x_13); -lean_dec(x_12); -lean_dec(x_11); -x_27 = !lean_is_exclusive(x_26); -if (x_27 == 0) +x_54 = lean_ctor_get(x_12, 2); +lean_inc(x_54); +x_55 = l_Lean_Elab_wfRecursion___lambda__6___closed__1; +x_56 = 1; +x_57 = l_Lean_Option_set___at_Lean_Elab_Term_withoutMacroStackAtErr___spec__1(x_54, x_55, x_56); +x_58 = l_Lean_Elab_wfRecursion___lambda__6___closed__2; +x_59 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_57, x_58); +x_60 = lean_st_ref_get(x_13, x_17); +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_ctor_get(x_61, 0); +lean_inc(x_63); +lean_dec(x_61); +x_64 = l_Lean_Kernel_isDiagnosticsEnabled(x_63); +lean_dec(x_63); +if (x_64 == 0) { -return x_26; -} -else +if (x_59 == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_26, 0); -x_29 = lean_ctor_get(x_26, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_26); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; -} +x_65 = x_56; +goto block_108; } else { -lean_object* x_31; lean_object* x_32; -x_31 = lean_box(0); -x_32 = l_Lean_Elab_wfRecursion___lambda__3(x_18, x_1, x_2, x_3, x_4, x_8, x_5, x_6, x_7, x_31, x_10, x_11, x_12, x_13, x_14, x_15, x_19); -lean_dec(x_18); -return x_32; +uint8_t x_109; +x_109 = 0; +x_65 = x_109; +goto block_108; } } else { -uint8_t x_33; -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_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_1); -x_33 = !lean_is_exclusive(x_17); -if (x_33 == 0) +if (x_59 == 0) { -return x_17; +uint8_t x_110; +x_110 = 0; +x_65 = x_110; +goto block_108; } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_34 = lean_ctor_get(x_17, 0); -x_35 = lean_ctor_get(x_17, 1); -lean_inc(x_35); -lean_inc(x_34); -lean_dec(x_17); -x_36 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_35); -return x_36; -} +x_65 = x_56; +goto block_108; } } -} -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__5(lean_object* x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { -_start: +block_53: { -lean_object* x_14; +lean_object* x_19; +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_1); -x_14 = l_Lean_Elab_addAndCompilePartialRec(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_14) == 0) +lean_inc(x_16); +x_19 = l_Lean_Elab_addAndCompilePartialRec(x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_15; lean_object* x_16; size_t x_17; lean_object* x_18; -x_15 = lean_ctor_get(x_14, 1); -lean_inc(x_15); -lean_dec(x_14); -x_16 = lean_array_get_size(x_1); -x_17 = lean_usize_of_nat(x_16); -lean_dec(x_16); +lean_object* x_20; lean_object* x_21; size_t x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = lean_array_get_size(x_16); +x_22 = lean_usize_of_nat(x_21); +lean_dec(x_21); +lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -x_18 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__2(x_17, x_2, x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_15); -if (lean_obj_tag(x_18) == 0) +x_23 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__2(x_22, x_2, x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_20); +if (lean_obj_tag(x_23) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); -lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_3, 3); -lean_inc(x_21); -lean_dec(x_3); +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_ctor_get(x_4, 3); +lean_inc(x_26); +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_19); -x_22 = l_Lean_Elab_WF_registerEqnsInfo(x_19, x_21, x_4, x_5, x_9, x_10, x_11, x_12, x_20); -if (lean_obj_tag(x_22) == 0) +lean_inc(x_24); +x_27 = l_Lean_Elab_WF_registerEqnsInfo(x_24, x_26, x_5, x_6, x_10, x_11, x_12, x_13, x_25); +if (lean_obj_tag(x_27) == 0) { -lean_object* x_23; lean_object* x_24; size_t x_25; lean_object* x_26; lean_object* x_27; -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = lean_array_get_size(x_19); -x_25 = lean_usize_of_nat(x_24); +lean_object* x_28; lean_object* x_29; size_t x_30; lean_object* x_31; lean_object* x_32; +x_28 = lean_ctor_get(x_27, 1); +lean_inc(x_28); +lean_dec(x_27); +x_29 = lean_array_get_size(x_24); +x_30 = lean_usize_of_nat(x_29); +lean_dec(x_29); +x_31 = lean_box(0); +x_32 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(x_2, x_24, x_30, x_2, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_28); lean_dec(x_24); -x_26 = lean_box(0); -x_27 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3(x_19, x_25, x_2, x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_23); -lean_dec(x_19); -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_32) == 0) { -uint8_t x_28; -x_28 = !lean_is_exclusive(x_27); -if (x_28 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) { -lean_object* x_29; -x_29 = lean_ctor_get(x_27, 0); -lean_dec(x_29); -lean_ctor_set(x_27, 0, x_26); -return x_27; +lean_object* x_34; +x_34 = lean_ctor_get(x_32, 0); +lean_dec(x_34); +lean_ctor_set(x_32, 0, x_31); +return x_32; } else { -lean_object* x_30; lean_object* x_31; -x_30 = lean_ctor_get(x_27, 1); -lean_inc(x_30); -lean_dec(x_27); -x_31 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_31, 0, x_26); -lean_ctor_set(x_31, 1, x_30); -return x_31; +lean_object* x_35; lean_object* x_36; +x_35 = lean_ctor_get(x_32, 1); +lean_inc(x_35); +lean_dec(x_32); +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_31); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } else { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_27); -if (x_32 == 0) +uint8_t x_37; +x_37 = !lean_is_exclusive(x_32); +if (x_37 == 0) { -return x_27; +return x_32; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_27, 0); -x_34 = lean_ctor_get(x_27, 1); -lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_27); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_34); -return x_35; +lean_object* x_38; lean_object* x_39; lean_object* x_40; +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_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 { -uint8_t x_36; -lean_dec(x_19); +uint8_t x_41; +lean_dec(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); -x_36 = !lean_is_exclusive(x_22); -if (x_36 == 0) +x_41 = !lean_is_exclusive(x_27); +if (x_41 == 0) { -return x_22; +return x_27; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_22, 0); -x_38 = lean_ctor_get(x_22, 1); -lean_inc(x_38); -lean_inc(x_37); -lean_dec(x_22); -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_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_27, 0); +x_43 = lean_ctor_get(x_27, 1); +lean_inc(x_43); +lean_inc(x_42); +lean_dec(x_27); +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_40; +uint8_t x_45; +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -x_40 = !lean_is_exclusive(x_18); -if (x_40 == 0) +x_45 = !lean_is_exclusive(x_23); +if (x_45 == 0) { -return x_18; +return x_23; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_18, 0); -x_42 = lean_ctor_get(x_18, 1); -lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_18); -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_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_23, 0); +x_47 = lean_ctor_get(x_23, 1); +lean_inc(x_47); +lean_inc(x_46); +lean_dec(x_23); +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_44; +uint8_t x_49; +lean_dec(x_16); +lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); -lean_dec(x_7); +lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_1); -x_44 = !lean_is_exclusive(x_14); -if (x_44 == 0) +x_49 = !lean_is_exclusive(x_19); +if (x_49 == 0) { -return x_14; +return x_19; } else { -lean_object* x_45; lean_object* x_46; lean_object* x_47; -x_45 = lean_ctor_get(x_14, 0); -x_46 = lean_ctor_get(x_14, 1); -lean_inc(x_46); -lean_inc(x_45); -lean_dec(x_14); -x_47 = lean_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_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_19, 0); +x_51 = lean_ctor_get(x_19, 1); +lean_inc(x_51); +lean_inc(x_50); +lean_dec(x_19); +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_wfRecursion___lambda__6(size_t x_1, size_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14) { -_start: -{ -lean_object* x_15; -lean_dec(x_7); -lean_inc(x_13); -lean_inc(x_12); -x_15 = l_Array_mapMUnsafe_map___at_Lean_Elab_addAndCompileUnsafe___spec__1(x_1, x_2, x_3, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_15) == 0) +block_108: { -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_4); -x_18 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_isOnlyOneUnaryDef(x_16, x_4, x_10, x_11, x_12, x_13, x_17); -if (lean_obj_tag(x_18) == 0) +if (x_65 == 0) { -lean_object* x_19; uint8_t x_20; -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_66; lean_object* x_67; lean_object* x_68; uint8_t x_69; +x_66 = lean_st_ref_take(x_13, x_62); +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_is_exclusive(x_67); +if (x_69 == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_21 = lean_ctor_get(x_18, 1); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_ctor_get(x_5, 3); -lean_inc(x_22); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_24, 0, x_22); -lean_ctor_set(x_24, 1, x_23); -x_25 = 0; -x_26 = lean_box(x_25); -lean_inc(x_5); -x_27 = lean_alloc_closure((void*)(l_Lean_Elab_addNonRec___boxed), 10, 3); -lean_closure_set(x_27, 0, x_5); -lean_closure_set(x_27, 1, x_26); -lean_closure_set(x_27, 2, x_24); +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_70 = lean_ctor_get(x_67, 0); +x_71 = lean_ctor_get(x_67, 4); +lean_dec(x_71); +x_72 = l_Lean_Kernel_enableDiag(x_70, x_59); +x_73 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4; +lean_ctor_set(x_67, 4, x_73); +lean_ctor_set(x_67, 0, x_72); +x_74 = lean_st_ref_set(x_13, x_67, x_68); +x_75 = lean_ctor_get(x_74, 1); +lean_inc(x_75); +lean_dec(x_74); +x_76 = lean_box(0); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); +lean_inc(x_6); lean_inc(x_9); lean_inc(x_8); -x_28 = l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__4(x_25, x_27, x_8, x_9, x_10, x_11, x_12, x_13, x_21); -if (lean_obj_tag(x_28) == 0) -{ -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); -lean_inc(x_13); -lean_inc(x_12); +lean_inc(x_4); lean_inc(x_11); lean_inc(x_10); -lean_inc(x_9); -lean_inc(x_8); lean_inc(x_5); lean_inc(x_16); -lean_inc(x_6); -lean_inc(x_4); -x_30 = l___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs(x_4, x_6, x_16, x_5, x_8, x_9, x_10, x_11, x_12, x_13, x_29); -if (lean_obj_tag(x_30) == 0) +x_77 = l_Lean_Elab_wfRecursion___lambda__5(x_57, x_59, x_16, x_5, x_10, x_11, x_4, x_8, x_9, x_6, x_76, x_12, x_13, x_75); +if (lean_obj_tag(x_77) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); -lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_Elab_wfRecursion___lambda__5(x_16, x_2, x_5, x_4, x_6, x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_32); -lean_dec(x_31); -return x_33; +lean_object* x_78; +x_78 = lean_ctor_get(x_77, 1); +lean_inc(x_78); +lean_dec(x_77); +x_18 = x_78; +goto block_53; } else { -uint8_t x_34; +uint8_t x_79; lean_dec(x_16); lean_dec(x_13); lean_dec(x_12); @@ -8199,29 +9185,80 @@ lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_34 = !lean_is_exclusive(x_30); -if (x_34 == 0) +x_79 = !lean_is_exclusive(x_77); +if (x_79 == 0) { -return x_30; +return x_77; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_30, 0); -x_36 = lean_ctor_get(x_30, 1); -lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_30); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_80 = lean_ctor_get(x_77, 0); +x_81 = lean_ctor_get(x_77, 1); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_77); +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; } } } else { -uint8_t x_38; +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +x_83 = lean_ctor_get(x_67, 0); +x_84 = lean_ctor_get(x_67, 1); +x_85 = lean_ctor_get(x_67, 2); +x_86 = lean_ctor_get(x_67, 3); +x_87 = lean_ctor_get(x_67, 5); +x_88 = lean_ctor_get(x_67, 6); +lean_inc(x_88); +lean_inc(x_87); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_dec(x_67); +x_89 = l_Lean_Kernel_enableDiag(x_83, x_59); +x_90 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4; +x_91 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_91, 0, x_89); +lean_ctor_set(x_91, 1, x_84); +lean_ctor_set(x_91, 2, x_85); +lean_ctor_set(x_91, 3, x_86); +lean_ctor_set(x_91, 4, x_90); +lean_ctor_set(x_91, 5, x_87); +lean_ctor_set(x_91, 6, x_88); +x_92 = lean_st_ref_set(x_13, x_91, x_68); +x_93 = lean_ctor_get(x_92, 1); +lean_inc(x_93); +lean_dec(x_92); +x_94 = lean_box(0); +lean_inc(x_13); +lean_inc(x_12); +lean_inc(x_6); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_11); +lean_inc(x_10); +lean_inc(x_5); +lean_inc(x_16); +x_95 = l_Lean_Elab_wfRecursion___lambda__5(x_57, x_59, x_16, x_5, x_10, x_11, x_4, x_8, x_9, x_6, x_94, x_12, x_13, x_93); +if (lean_obj_tag(x_95) == 0) +{ +lean_object* x_96; +x_96 = lean_ctor_get(x_95, 1); +lean_inc(x_96); +lean_dec(x_95); +x_18 = x_96; +goto block_53; +} +else +{ +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_dec(x_16); lean_dec(x_13); lean_dec(x_12); @@ -8232,62 +9269,56 @@ lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_38 = !lean_is_exclusive(x_28); -if (x_38 == 0) -{ -return x_28; +x_97 = lean_ctor_get(x_95, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_95, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_95)) { + lean_ctor_release(x_95, 0); + lean_ctor_release(x_95, 1); + x_99 = x_95; +} else { + lean_dec_ref(x_95); + x_99 = lean_box(0); } -else -{ -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; +if (lean_is_scalar(x_99)) { + x_100 = lean_alloc_ctor(1, 2, 0); +} else { + x_100 = x_99; +} +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_98); +return x_100; } } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; uint8_t x_46; lean_object* x_47; -x_42 = lean_ctor_get(x_18, 1); -lean_inc(x_42); -lean_dec(x_18); -x_43 = lean_ctor_get(x_5, 3); -lean_inc(x_43); -x_44 = lean_box(0); -x_45 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_45, 0, x_43); -lean_ctor_set(x_45, 1, x_44); -x_46 = 0; +lean_object* x_101; lean_object* x_102; +x_101 = lean_box(0); lean_inc(x_13); lean_inc(x_12); -lean_inc(x_11); -lean_inc(x_10); +lean_inc(x_6); lean_inc(x_9); lean_inc(x_8); +lean_inc(x_4); +lean_inc(x_11); +lean_inc(x_10); lean_inc(x_5); -x_47 = l___private_Lean_Elab_PreDefinition_Basic_0__Lean_Elab_addNonRecAux(x_5, x_46, x_45, x_46, x_8, x_9, x_10, x_11, x_12, x_13, x_42); -if (lean_obj_tag(x_47) == 0) +lean_inc(x_16); +x_102 = l_Lean_Elab_wfRecursion___lambda__5(x_57, x_59, x_16, x_5, x_10, x_11, x_4, x_8, x_9, x_6, x_101, x_12, x_13, x_62); +if (lean_obj_tag(x_102) == 0) { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -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_wfRecursion___lambda__5(x_16, x_2, x_5, x_4, x_6, x_48, x_8, x_9, x_10, x_11, x_12, x_13, x_49); -lean_dec(x_48); -return x_50; +lean_object* x_103; +x_103 = lean_ctor_get(x_102, 1); +lean_inc(x_103); +lean_dec(x_102); +x_18 = x_103; +goto block_53; } else { -uint8_t x_51; +uint8_t x_104; lean_dec(x_16); lean_dec(x_13); lean_dec(x_12); @@ -8298,63 +9329,31 @@ lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_51 = !lean_is_exclusive(x_47); -if (x_51 == 0) +x_104 = !lean_is_exclusive(x_102); +if (x_104 == 0) { -return x_47; +return x_102; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_47, 0); -x_53 = lean_ctor_get(x_47, 1); -lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_47); -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_105; lean_object* x_106; lean_object* x_107; +x_105 = lean_ctor_get(x_102, 0); +x_106 = lean_ctor_get(x_102, 1); +lean_inc(x_106); +lean_inc(x_105); +lean_dec(x_102); +x_107 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_107, 0, x_105); +lean_ctor_set(x_107, 1, x_106); +return x_107; } -else -{ -uint8_t x_55; -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_6); -lean_dec(x_5); -lean_dec(x_4); -x_55 = !lean_is_exclusive(x_18); -if (x_55 == 0) -{ -return x_18; } -else -{ -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_18, 0); -x_57 = lean_ctor_get(x_18, 1); -lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_18); -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_59; +uint8_t x_111; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -8364,23 +9363,23 @@ lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_59 = !lean_is_exclusive(x_15); -if (x_59 == 0) +x_111 = !lean_is_exclusive(x_15); +if (x_111 == 0) { return x_15; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_15, 0); -x_61 = lean_ctor_get(x_15, 1); -lean_inc(x_61); -lean_inc(x_60); +lean_object* x_112; lean_object* x_113; lean_object* x_114; +x_112 = lean_ctor_get(x_15, 0); +x_113 = lean_ctor_get(x_15, 1); +lean_inc(x_113); +lean_inc(x_112); lean_dec(x_15); -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; +x_114 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_114, 0, x_112); +lean_ctor_set(x_114, 1, x_113); +return x_114; } } } @@ -8469,7 +9468,7 @@ 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_wfRecursion___lambda__6(x_3, x_4, x_5, x_2, x_21, x_6, x_28, x_8, x_9, x_10, x_11, x_12, x_13, x_27); +x_29 = l_Lean_Elab_wfRecursion___lambda__6(x_3, x_4, x_5, x_21, x_2, x_6, x_28, x_8, x_9, x_10, x_11, x_12, x_13, x_27); return x_29; } else @@ -8505,7 +9504,8 @@ 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_wfRecursion___lambda__6(x_3, x_4, x_5, x_2, x_21, x_6, x_43, x_8, x_9, x_10, x_11, x_12, x_13, x_44); +x_45 = l_Lean_Elab_wfRecursion___lambda__6(x_3, x_4, x_5, x_21, x_2, x_6, x_43, x_8, x_9, x_10, x_11, x_12, x_13, x_44); +lean_dec(x_43); return x_45; } } @@ -8552,7 +9552,7 @@ lean_inc(x_9); lean_inc(x_8); lean_inc(x_4); lean_inc(x_1); -x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__10(x_1, x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_13 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__12(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; @@ -8565,7 +9565,7 @@ lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); -x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11(x_2, x_3, x_4, x_6, x_7, x_8, x_9, x_10, x_11, x_15); +x_16 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13(x_2, x_3, x_4, 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; @@ -8767,7 +9767,7 @@ lean_dec(x_66); x_69 = lean_box(0); lean_inc(x_7); lean_inc(x_6); -x_70 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__9(x_13, x_17, x_11, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_67); +x_70 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__11(x_13, x_17, x_11, x_69, x_2, x_3, x_4, x_5, x_6, x_7, x_67); if (lean_obj_tag(x_70) == 0) { lean_object* x_71; lean_object* x_72; @@ -9047,7 +10047,7 @@ x_23 = lean_ctor_get(x_20, 1); lean_inc(x_23); lean_dec(x_20); x_24 = l_Lean_Elab_wfRecursion___closed__1; -x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(x_13, x_17, x_11, x_24); +x_25 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__7(x_13, x_17, x_11, x_24); x_26 = lean_ctor_get(x_25, 0); lean_inc(x_26); x_27 = lean_ctor_get(x_25, 1); @@ -9123,7 +10123,7 @@ lean_dec(x_15); x_38 = lean_usize_of_nat(x_30); lean_dec(x_30); x_39 = lean_box(0); -x_40 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6(x_26, x_27, x_11, x_38, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_19); +x_40 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8(x_26, x_27, x_11, x_38, x_39, x_2, x_3, x_4, x_5, x_6, x_7, x_19); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); @@ -9169,7 +10169,7 @@ lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_48 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8(x_21, x_22, x_26, x_26, x_45, x_47, lean_box(0), x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_19); +x_48 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10(x_21, x_22, x_26, x_26, x_45, x_47, lean_box(0), x_46, x_2, x_3, x_4, x_5, x_6, x_7, x_19); lean_dec(x_26); if (lean_obj_tag(x_48) == 0) { @@ -9337,30 +10337,62 @@ lean_dec(x_4); return x_13; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_EXPORT lean_object* l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -size_t x_12; size_t x_13; lean_object* x_14; -x_12 = lean_unbox_usize(x_2); +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); lean_dec(x_2); -x_13 = lean_unbox_usize(x_3); +x_11 = l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___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_14 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_6 = l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4(x_1, x_4, x_5); lean_dec(x_1); -return x_14; +x_7 = lean_box(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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: +{ +size_t x_13; size_t x_14; size_t x_15; lean_object* x_16; +x_13 = lean_unbox_usize(x_1); +lean_dec(x_1); +x_14 = lean_unbox_usize(x_3); +lean_dec(x_3); +x_15 = lean_unbox_usize(x_4); +lean_dec(x_4); +x_16 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(x_13, x_2, x_14, x_15, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +lean_dec(x_2); +return x_16; } } -LEAN_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___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_EXPORT lean_object* l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___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) { _start: { uint8_t x_10; lean_object* x_11; x_10 = lean_unbox(x_1); lean_dec(x_1); -x_11 = l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__4(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_Lean_Elab_withEnableInfoTree___at_Lean_Elab_wfRecursion___spec__6(x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); return x_11; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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; @@ -9368,12 +10400,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_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5(x_1, x_5, x_6, x_4); +x_7 = l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__7(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_Elab_wfRecursion___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___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) { _start: { size_t x_13; size_t x_14; lean_object* x_15; @@ -9381,7 +10413,7 @@ x_13 = lean_unbox_usize(x_3); lean_dec(x_3); x_14 = lean_unbox_usize(x_4); lean_dec(x_4); -x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +x_15 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8(x_1, x_2, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -9393,11 +10425,11 @@ lean_dec(x_1); return x_15; } } -LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13, lean_object* x_14, lean_object* x_15) { +LEAN_EXPORT lean_object* l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___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, lean_object* x_14, lean_object* x_15) { _start: { lean_object* x_16; -x_16 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); +x_16 = l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14, x_15); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); @@ -9405,7 +10437,7 @@ lean_dec(x_1); return x_16; } } -LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___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_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__11___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -9413,7 +10445,7 @@ 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_Lean_Elab_wfRecursion___spec__9(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_Lean_Elab_wfRecursion___spec__11(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_8); lean_dec(x_7); lean_dec(x_6); @@ -9422,7 +10454,7 @@ lean_dec(x_1); return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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_object* x_11) { _start: { size_t x_12; size_t x_13; lean_object* x_14; @@ -9430,13 +10462,13 @@ x_12 = lean_unbox_usize(x_2); lean_dec(x_2); x_13 = lean_unbox_usize(x_3); lean_dec(x_3); -x_14 = l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__10(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_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__12(x_1, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11); lean_dec(x_6); lean_dec(x_5); return x_14; } } -LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___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_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { size_t x_11; size_t x_12; lean_object* x_13; @@ -9444,7 +10476,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_Lean_Elab_wfRecursion___spec__11(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_Lean_Elab_wfRecursion___spec__13(x_11, x_12, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); lean_dec(x_4); return x_13; @@ -9535,15 +10567,15 @@ x_19 = l_Lean_Elab_wfRecursion___lambda__4(x_1, x_17, x_18, x_4, x_5, x_6, x_7, return x_19; } } -LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___lambda__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12, lean_object* x_13) { +LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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) { _start: { -size_t x_14; lean_object* x_15; -x_14 = lean_unbox_usize(x_2); +uint8_t x_15; lean_object* x_16; +x_15 = lean_unbox(x_2); lean_dec(x_2); -x_15 = l_Lean_Elab_wfRecursion___lambda__5(x_1, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -lean_dec(x_6); -return x_15; +x_16 = l_Lean_Elab_wfRecursion___lambda__5(x_1, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_11); +return x_16; } } LEAN_EXPORT lean_object* l_Lean_Elab_wfRecursion___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) { @@ -9555,6 +10587,7 @@ lean_dec(x_1); x_16 = lean_unbox_usize(x_2); lean_dec(x_2); x_17 = l_Lean_Elab_wfRecursion___lambda__6(x_15, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +lean_dec(x_7); return x_17; } } @@ -9585,7 +10618,7 @@ lean_dec(x_5); return x_15; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1() { _start: { lean_object* x_1; @@ -9593,27 +10626,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1; 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_PreDefinition_WF_Main___hyg_2773____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2; x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__1; 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_PreDefinition_WF_Main___hyg_2773____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__4() { _start: { lean_object* x_1; @@ -9621,17 +10654,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__3; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__4; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__4; 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_PreDefinition_WF_Main___hyg_2773____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__6() { _start: { lean_object* x_1; @@ -9639,37 +10672,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__6; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____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_PreDefinition_WF_Main___hyg_2773____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__7; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1; 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_PreDefinition_WF_Main___hyg_2773____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__8; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__8; x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__1; 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_PreDefinition_WF_Main___hyg_2773____closed__10() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10() { _start: { lean_object* x_1; @@ -9677,17 +10710,17 @@ x_1 = lean_mk_string_from_bytes("PreDefinition", 13); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__9; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__10; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10; 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_PreDefinition_WF_Main___hyg_2773____closed__12() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12() { _start: { lean_object* x_1; @@ -9695,17 +10728,17 @@ x_1 = lean_mk_string_from_bytes("WF", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__12; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12; 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_PreDefinition_WF_Main___hyg_2773____closed__14() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14() { _start: { lean_object* x_1; @@ -9713,17 +10746,17 @@ x_1 = lean_mk_string_from_bytes("Main", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__14; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14; 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_PreDefinition_WF_Main___hyg_2773____closed__16() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16() { _start: { lean_object* x_1; @@ -9731,33 +10764,33 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__16; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16; 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_PreDefinition_WF_Main___hyg_2773____closed__18() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17; -x_2 = lean_unsigned_to_nat(2773u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17; +x_2 = lean_unsigned_to_nat(2830u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830_(lean_object* x_1) { _start: { lean_object* x_2; uint8_t x_3; lean_object* x_4; lean_object* x_5; x_2 = l_Std_Range_forIn_loop___at___private_Lean_Elab_PreDefinition_WF_Main_0__Lean_Elab_addNonRecPreDefs___spec__1___closed__4; x_3 = 0; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__18; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -9889,32 +10922,78 @@ l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__1 = _init_ lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__1); l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__2(); lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__1___closed__2); -l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1(); -lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__3___closed__1); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__1); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__2); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__3); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__4); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__5); -l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6(); -lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__6___closed__6); -l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1(); -lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__1); -l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2(); -lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__2); -l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3(); -lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__3); -l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4(); -lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__8___closed__4); -l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__1); -l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2(); -lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__11___closed__2); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__1 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__1(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__1); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__2); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__3); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__4); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__5); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__6 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__6(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__6); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__7); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__8); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__9); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__10 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__10(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__10); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__11); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__12); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__13 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__13(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__13); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__14); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__15 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__15(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__15); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__16 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__16(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__16); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__17); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__18); +l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19 = _init_l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19(); +lean_mark_persistent(l_Lean_setReducibilityStatus___at_Lean_Elab_wfRecursion___spec__3___closed__19); +l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__1); +l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__2); +l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__3); +l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4 = _init_l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4(); +lean_mark_persistent(l_Array_anyMUnsafe_any___at_Lean_Elab_wfRecursion___spec__4___closed__4); +l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1(); +lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_wfRecursion___spec__5___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__1); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__2); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__3); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__4); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__5); +l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6 = _init_l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6(); +lean_mark_persistent(l_Array_foldlMUnsafe_fold___at_Lean_Elab_wfRecursion___spec__8___closed__6); +l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__1); +l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__2); +l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__3); +l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4 = _init_l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4(); +lean_mark_persistent(l_Array_mapIdxM_map___at_Lean_Elab_wfRecursion___spec__10___closed__4); +l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__1); +l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2 = _init_l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2(); +lean_mark_persistent(l_Array_mapMUnsafe_map___at_Lean_Elab_wfRecursion___spec__13___closed__2); l_Lean_Elab_wfRecursion___lambda__1___closed__1 = _init_l_Lean_Elab_wfRecursion___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__1___closed__1); l_Lean_Elab_wfRecursion___lambda__1___closed__2 = _init_l_Lean_Elab_wfRecursion___lambda__1___closed__2(); @@ -9927,6 +11006,12 @@ l_Lean_Elab_wfRecursion___lambda__4___closed__1 = _init_l_Lean_Elab_wfRecursion_ lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__4___closed__1); l_Lean_Elab_wfRecursion___lambda__4___closed__2 = _init_l_Lean_Elab_wfRecursion___lambda__4___closed__2(); lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__4___closed__2); +l_Lean_Elab_wfRecursion___lambda__5___closed__1 = _init_l_Lean_Elab_wfRecursion___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__5___closed__1); +l_Lean_Elab_wfRecursion___lambda__6___closed__1 = _init_l_Lean_Elab_wfRecursion___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__6___closed__1); +l_Lean_Elab_wfRecursion___lambda__6___closed__2 = _init_l_Lean_Elab_wfRecursion___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__6___closed__2); l_Lean_Elab_wfRecursion___lambda__7___closed__1 = _init_l_Lean_Elab_wfRecursion___lambda__7___closed__1(); lean_mark_persistent(l_Lean_Elab_wfRecursion___lambda__7___closed__1); l_Lean_Elab_wfRecursion___lambda__7___closed__2 = _init_l_Lean_Elab_wfRecursion___lambda__7___closed__2(); @@ -9941,43 +11026,43 @@ l_Lean_Elab_wfRecursion___closed__2 = _init_l_Lean_Elab_wfRecursion___closed__2( lean_mark_persistent(l_Lean_Elab_wfRecursion___closed__2); l_Lean_Elab_wfRecursion___closed__3 = _init_l_Lean_Elab_wfRecursion___closed__3(); lean_mark_persistent(l_Lean_Elab_wfRecursion___closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__9); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__10(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__10); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__11); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__12 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__12(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__12); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__13); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__14 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__14(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__14); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__15); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__16 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__16(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__16); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__17); -l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__18 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__18(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773____closed__18); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2773_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__9); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__10); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__11); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__12); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__13); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__14); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__15); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__16); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__17); +l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830____closed__18); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_PreDefinition_WF_Main___hyg_2830_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationArgument.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationArgument.c index bbe7487a9e13..6f6a9e1ffe53 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationArgument.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationArgument.c @@ -78,7 +78,7 @@ lean_object* l_Lean_Syntax_node2(lean_object*, lean_object*, lean_object*, lean_ static lean_object* l_Lean_Elab_WF_TerminationArgument_elab___closed__9; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Delaborator_SubExpr_descend___at_Lean_Elab_WF_TerminationArgument_delab_go___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationArgument_elab___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* l_Array_append___rarg(lean_object*, lean_object*); lean_object* l_Lean_SubExpr_Pos_push(lean_object*, lean_object*); @@ -353,14 +353,14 @@ lean_closure_set(x_19, 1, x_15); lean_closure_set(x_19, 2, x_17); lean_closure_set(x_19, 3, x_18); lean_closure_set(x_19, 4, x_15); -x_20 = 0; +x_20 = 1; 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_21 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_19, x_20, x_16, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +x_21 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_19, x_20, x_7, x_8, x_9, x_10, x_11, x_12, x_13); if (lean_obj_tag(x_21) == 0) { lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; diff --git a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c index d8c8fb1421af..b2e450ea7d7c 100644 --- a/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c +++ b/stage0/stdlib/Lean/Elab/PreDefinition/WF/TerminationHint.c @@ -53,6 +53,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_WF_TerminationHints_rememberExtraParams___b static lean_object* l_Lean_Elab_WF_instInhabitedTerminationBy___closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_WF_TerminationBy_checkVars___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__10; +lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__5; static lean_object* l_Lean_Elab_WF_instInhabitedTerminationBy___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_WF_instInhabitedDecreasingBy; @@ -98,7 +99,6 @@ static lean_object* l_Lean_Elab_WF_TerminationHints_ensureNone___closed__6; lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_WF_TerminationHints_isNotNone(lean_object*); static lean_object* l_Lean_Elab_WF_TerminationBy_checkVars___closed__5; -lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__3___closed__6; LEAN_EXPORT uint8_t l_Lean_Elab_WF_TerminationBy_synthetic___default; LEAN_EXPORT lean_object* l_Lean_Elab_WF_elabTerminationHints___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -334,7 +334,7 @@ x_17 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_17, 0, x_15); lean_ctor_set(x_17, 1, x_16); x_18 = 2; -x_19 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_12, x_17, x_18, x_3, x_4, x_5); +x_19 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_12, x_17, x_18, x_3, x_4, x_5); return x_19; } } @@ -357,7 +357,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 = 2; -x_29 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_22, x_27, x_28, x_3, x_4, x_5); +x_29 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_22, x_27, x_28, x_3, x_4, x_5); return x_29; } else @@ -374,7 +374,7 @@ x_35 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_35, 0, x_33); lean_ctor_set(x_35, 1, x_34); x_36 = 2; -x_37 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_30, x_35, x_36, x_3, x_4, x_5); +x_37 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_30, x_35, x_36, x_3, x_4, x_5); return x_37; } } @@ -401,7 +401,7 @@ x_45 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_45, 0, x_43); lean_ctor_set(x_45, 1, x_44); x_46 = 2; -x_47 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_40, x_45, x_46, x_3, x_4, x_5); +x_47 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_40, x_45, x_46, x_3, x_4, x_5); return x_47; } else @@ -418,7 +418,7 @@ x_53 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_53, 0, x_51); lean_ctor_set(x_53, 1, x_52); x_54 = 2; -x_55 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_48, x_53, x_54, x_3, x_4, x_5); +x_55 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_48, x_53, x_54, x_3, x_4, x_5); return x_55; } } @@ -436,7 +436,7 @@ x_61 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_61, 0, x_59); lean_ctor_set(x_61, 1, x_60); x_62 = 2; -x_63 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_56, x_61, x_62, x_3, x_4, x_5); +x_63 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_56, x_61, x_62, x_3, x_4, x_5); return x_63; } } diff --git a/stage0/stdlib/Lean/Elab/Print.c b/stage0/stdlib/Lean/Elab/Print.c index 5f496d63723d..5a3b35eca2ac 100644 --- a/stage0/stdlib/Lean/Elab/Print.c +++ b/stage0/stdlib/Lean/Elab/Print.c @@ -32,6 +32,7 @@ static lean_object* l_Lean_Elab_Command_elabPrint___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_private_to_user_name(lean_object*); lean_object* l_Lean_logAt___at_Lean_Elab_Command_elabCommand___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_Elab_Command_0__Lean_Elab_Command_elabCommandUsing___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printStructure___closed__1; @@ -109,7 +110,6 @@ static lean_object* l_panic___at___private_Lean_Elab_Print_0__Lean_Elab_Command_ static lean_object* l___regBuiltin_Lean_Elab_Command_elabPrint___closed__1; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__25; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__3; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Print_0__Lean_Elab_Command_printEqnsOf___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_mkHeader___closed__14; static lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printInduct___closed__5; @@ -1800,7 +1800,7 @@ x_40 = l_Lean_Name_toString(x_16, x_39); x_41 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_41, 0, x_40); lean_inc(x_7); -x_42 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +x_42 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_26); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; @@ -1962,7 +1962,7 @@ x_76 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_76, 0, x_75); lean_ctor_set(x_76, 1, x_74); lean_inc(x_7); -x_77 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_26); +x_77 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_37, x_7, x_8, x_9, x_10, x_11, x_12, x_26); if (lean_obj_tag(x_77) == 0) { lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; @@ -2299,7 +2299,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Print_0__Lean_Elab_Command_printS lean_object* x_10; lean_inc(x_3); lean_inc(x_1); -x_10 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_10 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_10) == 0) { lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; diff --git a/stage0/stdlib/Lean/Elab/StructInst.c b/stage0/stdlib/Lean/Elab/StructInst.c index 2c8a68019be6..998ae11cde00 100644 --- a/stage0/stdlib/Lean/Elab/StructInst.c +++ b/stage0/stdlib/Lean/Elab/StructInst.c @@ -68,6 +68,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_StructInst_e static lean_object* l_List_mapM_loop___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_expandParentFields___spec__3___closed__8; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_findField_x3f___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_defaultMissing_x3f___boxed(lean_object*); +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkCtorHeaderAux___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_addMissingFields___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_isModifyOp_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -338,7 +339,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_Stru 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_AssocList_find_x3f___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__5(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_StructInst_instToFormatFieldLHS___closed__3; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_mkHole(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_instToFormatStruct; static lean_object* l_Lean_Elab_Term_StructInst_expandStructInstFieldAbbrev___lambda__1___closed__7; @@ -509,7 +509,7 @@ LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_StructInst lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_mkFieldMap___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_format___at_Lean_Elab_Term_StructInst_formatStruct___spec__3___closed__5; -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, 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_Elab_Term_StructInst_formatStruct___spec__2___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_Elab_Term_StructInst_DefaultFields_allDefaultMissing___at_Lean_Elab_Term_StructInst_DefaultFields_propagateLoop___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -23887,7 +23887,7 @@ x_31 = lean_ctor_get(x_27, 0); lean_inc(x_31); lean_dec(x_27); lean_inc(x_8); -x_32 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_24); +x_32 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_31, x_8, x_9, x_10, x_11, x_12, x_13, x_24); if (lean_obj_tag(x_32) == 0) { lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; @@ -28325,14 +28325,14 @@ lean_dec(x_4); x_12 = lean_alloc_closure((void*)(l___private_Lean_Elab_StructInst_0__Lean_Elab_Term_StructInst_elabStruct), 9, 2); lean_closure_set(x_12, 0, x_1); lean_closure_set(x_12, 1, x_2); -x_13 = 1; +x_13 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; diff --git a/stage0/stdlib/Lean/Elab/Structure.c b/stage0/stdlib/Lean/Elab/Structure.c index 9d96d1e34f40..947185a2a957 100644 --- a/stage0/stdlib/Lean/Elab/Structure.c +++ b/stage0/stdlib/Lean/Elab/Structure.c @@ -72,6 +72,7 @@ lean_object* l_Lean_Elab_Command_checkResultingUniverse(lean_object*, lean_objec lean_object* l_Lean_getProjectionFnInfo_x3f___at_Lean_Meta_getStuckMVar_x3f___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_mkCoercionToCopiedParent(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__26; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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*); lean_object* l_Lean_Elab_sortDeclLevelParams(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___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*, lean_object*, lean_object*); @@ -378,7 +379,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_co LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_addProjections(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___rarg___lambda__4___closed__3; -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_registerStructure___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Command_checkValidFieldModifier___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -536,7 +536,6 @@ static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_ static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_copyNewFieldsFrom_copyFields_copy___rarg___closed__6; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getFieldType___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_Structure_0__Lean_Elab_Command_getFieldType___lambda__2___closed__2; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_withFields_go___spec__5(lean_object*); @@ -924,6 +923,7 @@ lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Elab_Command_elabStructure___closed__6; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_collectUniversesFromFields_go___spec__2___closed__1; +lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_addDeclarationRanges___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__20___closed__5; static lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_reprStructFieldKind____x40_Lean_Elab_Structure___hyg_337____closed__23; static lean_object* l_Lean_Elab_Command_checkValidCtorModifier___at___private_Lean_Elab_Structure_0__Lean_Elab_Command_expandCtor___spec__17___lambda__3___closed__2; @@ -11599,7 +11599,7 @@ x_20 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_20, 0, x_19); x_21 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_21, 0, x_20); -x_22 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_22 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_21, x_3, x_4, x_5, x_6, x_7, x_8, x_9); x_23 = !lean_is_exclusive(x_22); if (x_23 == 0) { @@ -12028,7 +12028,7 @@ x_19 = lean_ctor_get(x_17, 0); lean_inc(x_19); lean_dec(x_17); lean_inc(x_5); -x_20 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_15); +x_20 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_15); if (lean_obj_tag(x_20) == 0) { lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; @@ -12164,7 +12164,7 @@ x_44 = lean_ctor_get(x_41, 0); lean_inc(x_44); lean_dec(x_41); lean_inc(x_5); -x_45 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_39); +x_45 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_44, x_5, x_6, x_7, x_8, x_9, x_10, x_39); if (lean_obj_tag(x_45) == 0) { lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; @@ -14374,7 +14374,7 @@ 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); -x_69 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_68, x_6, x_7, x_8, x_9, x_10, x_11, x_51); +x_69 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_68, x_6, x_7, x_8, x_9, x_10, x_11, x_51); x_70 = lean_ctor_get(x_69, 0); lean_inc(x_70); x_71 = lean_ctor_get(x_69, 1); @@ -14642,7 +14642,7 @@ x_141 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_141, 0, x_140); x_142 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_142, 0, x_141); -x_143 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_142, x_6, x_7, x_8, x_9, x_95, x_11, x_125); +x_143 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_142, x_6, x_7, x_8, x_9, x_95, x_11, x_125); x_144 = lean_ctor_get(x_143, 0); lean_inc(x_144); x_145 = lean_ctor_get(x_143, 1); @@ -14833,153 +14833,154 @@ uint8_t x_14; x_14 = !lean_is_exclusive(x_11); if (x_14 == 0) { -lean_object* x_15; uint8_t x_16; lean_object* x_17; +lean_object* x_15; uint8_t x_16; uint8_t x_17; lean_object* x_18; x_15 = lean_ctor_get(x_11, 0); -x_16 = 0; +x_16 = 1; +x_17 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_16, x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_17) == 0) +x_18 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_16, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_17, 1); -lean_inc(x_18); -lean_dec(x_17); +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_18, 1); +lean_inc(x_19); +lean_dec(x_18); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -x_19 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_18); -if (lean_obj_tag(x_19) == 0) +x_20 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); +if (lean_obj_tag(x_20) == 0) { -lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_20 = lean_ctor_get(x_19, 0); -lean_inc(x_20); -x_21 = lean_ctor_get(x_19, 1); +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; +x_21 = lean_ctor_get(x_20, 0); lean_inc(x_21); -lean_dec(x_19); -x_22 = lean_box(0); -x_23 = 1; -x_24 = lean_box(x_23); -x_25 = lean_box(x_23); -x_26 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_26, 0, x_15); -lean_closure_set(x_26, 1, x_22); -lean_closure_set(x_26, 2, x_24); -lean_closure_set(x_26, 3, x_25); +x_22 = lean_ctor_get(x_20, 1); +lean_inc(x_22); +lean_dec(x_20); +x_23 = lean_box(0); +x_24 = 1; +x_25 = lean_box(x_24); +x_26 = lean_box(x_24); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_27, 0, x_15); +lean_closure_set(x_27, 1, x_23); +lean_closure_set(x_27, 2, x_25); +lean_closure_set(x_27, 3, x_26); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_27 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_26, x_3, x_4, x_5, x_6, x_7, x_8, x_21); -if (lean_obj_tag(x_27) == 0) +x_28 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_22); +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); -x_30 = 1; -x_31 = l_Lean_Meta_mkLambdaFVars(x_20, x_28, x_16, x_23, x_30, x_5, x_6, x_7, x_8, x_29); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); +x_31 = 1; +x_32 = l_Lean_Meta_mkLambdaFVars(x_21, x_29, x_17, x_24, x_31, x_5, x_6, x_7, x_8, x_30); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_31) == 0) +if (lean_obj_tag(x_32) == 0) { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_32); +if (x_33 == 0) { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -lean_ctor_set(x_11, 0, x_33); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_22); -lean_ctor_set(x_34, 1, x_11); -lean_ctor_set(x_31, 0, x_34); -return x_31; +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_32, 0); +lean_ctor_set(x_11, 0, x_34); +x_35 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_35, 0, x_23); +lean_ctor_set(x_35, 1, x_11); +lean_ctor_set(x_32, 0, x_35); +return x_32; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_35 = lean_ctor_get(x_31, 0); -x_36 = lean_ctor_get(x_31, 1); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_32, 0); +x_37 = lean_ctor_get(x_32, 1); +lean_inc(x_37); lean_inc(x_36); -lean_inc(x_35); -lean_dec(x_31); -lean_ctor_set(x_11, 0, x_35); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_22); -lean_ctor_set(x_37, 1, x_11); +lean_dec(x_32); +lean_ctor_set(x_11, 0, x_36); x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_37); -lean_ctor_set(x_38, 1, x_36); -return x_38; +lean_ctor_set(x_38, 0, x_23); +lean_ctor_set(x_38, 1, x_11); +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 { -uint8_t x_39; +uint8_t x_40; lean_free_object(x_11); -x_39 = !lean_is_exclusive(x_31); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_32); +if (x_40 == 0) { -return x_31; +return x_32; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_31, 0); -x_41 = lean_ctor_get(x_31, 1); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_32, 0); +x_42 = lean_ctor_get(x_32, 1); +lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_31); -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_32); +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 { -uint8_t x_43; -lean_dec(x_20); +uint8_t x_44; +lean_dec(x_21); lean_free_object(x_11); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_43 = !lean_is_exclusive(x_27); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_28); +if (x_44 == 0) { -return x_27; +return x_28; } 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_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_28, 0); +x_46 = lean_ctor_get(x_28, 1); +lean_inc(x_46); 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; +lean_dec(x_28); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -uint8_t x_47; +uint8_t x_48; lean_free_object(x_11); lean_dec(x_15); lean_dec(x_8); @@ -14988,29 +14989,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_47 = !lean_is_exclusive(x_19); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_20); +if (x_48 == 0) { -return x_19; +return x_20; } else { -lean_object* x_48; lean_object* x_49; lean_object* x_50; -x_48 = lean_ctor_get(x_19, 0); -x_49 = lean_ctor_get(x_19, 1); +lean_object* x_49; lean_object* x_50; lean_object* x_51; +x_49 = lean_ctor_get(x_20, 0); +x_50 = lean_ctor_get(x_20, 1); +lean_inc(x_50); lean_inc(x_49); -lean_inc(x_48); -lean_dec(x_19); -x_50 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_50, 0, x_48); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_dec(x_20); +x_51 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_51, 0, x_49); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -uint8_t x_51; +uint8_t x_52; lean_free_object(x_11); lean_dec(x_15); lean_dec(x_8); @@ -15020,208 +15021,209 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_51 = !lean_is_exclusive(x_17); -if (x_51 == 0) +x_52 = !lean_is_exclusive(x_18); +if (x_52 == 0) { -return x_17; +return x_18; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_17, 0); -x_53 = lean_ctor_get(x_17, 1); +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_18, 0); +x_54 = lean_ctor_get(x_18, 1); +lean_inc(x_54); lean_inc(x_53); -lean_inc(x_52); -lean_dec(x_17); -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_dec(x_18); +x_55 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_55, 0, x_53); +lean_ctor_set(x_55, 1, x_54); +return x_55; } } } else { -lean_object* x_55; uint8_t x_56; lean_object* x_57; -x_55 = lean_ctor_get(x_11, 0); -lean_inc(x_55); +lean_object* x_56; uint8_t x_57; uint8_t x_58; lean_object* x_59; +x_56 = lean_ctor_get(x_11, 0); +lean_inc(x_56); lean_dec(x_11); -x_56 = 0; +x_57 = 1; +x_58 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_57 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_56, x_56, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_57) == 0) +x_59 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_57, x_58, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_59) == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_57, 1); -lean_inc(x_58); -lean_dec(x_57); +lean_object* x_60; lean_object* x_61; +x_60 = lean_ctor_get(x_59, 1); +lean_inc(x_60); +lean_dec(x_59); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -x_59 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_58); -if (lean_obj_tag(x_59) == 0) +x_61 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_60); +if (lean_obj_tag(x_61) == 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_object* x_67; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_ctor_get(x_59, 1); -lean_inc(x_61); -lean_dec(x_59); -x_62 = lean_box(0); -x_63 = 1; -x_64 = lean_box(x_63); -x_65 = lean_box(x_63); -x_66 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); -lean_closure_set(x_66, 0, x_55); -lean_closure_set(x_66, 1, x_62); -lean_closure_set(x_66, 2, x_64); -lean_closure_set(x_66, 3, x_65); +lean_object* x_62; lean_object* x_63; lean_object* x_64; uint8_t x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; +x_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_box(0); +x_65 = 1; +x_66 = lean_box(x_65); +x_67 = lean_box(x_65); +x_68 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTerm___boxed), 11, 4); +lean_closure_set(x_68, 0, x_56); +lean_closure_set(x_68, 1, x_64); +lean_closure_set(x_68, 2, x_66); +lean_closure_set(x_68, 3, x_67); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_67 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_66, x_3, x_4, x_5, x_6, x_7, x_8, x_61); -if (lean_obj_tag(x_67) == 0) +x_69 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_68, x_3, x_4, x_5, x_6, x_7, x_8, x_63); +if (lean_obj_tag(x_69) == 0) { -lean_object* x_68; lean_object* x_69; uint8_t x_70; lean_object* x_71; -x_68 = lean_ctor_get(x_67, 0); -lean_inc(x_68); -x_69 = lean_ctor_get(x_67, 1); -lean_inc(x_69); -lean_dec(x_67); -x_70 = 1; -x_71 = l_Lean_Meta_mkLambdaFVars(x_60, x_68, x_56, x_63, x_70, x_5, x_6, x_7, x_8, x_69); +lean_object* x_70; lean_object* x_71; uint8_t x_72; lean_object* x_73; +x_70 = lean_ctor_get(x_69, 0); +lean_inc(x_70); +x_71 = lean_ctor_get(x_69, 1); +lean_inc(x_71); +lean_dec(x_69); +x_72 = 1; +x_73 = l_Lean_Meta_mkLambdaFVars(x_62, x_70, x_58, x_65, x_72, x_5, x_6, x_7, x_8, x_71); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_71) == 0) +if (lean_obj_tag(x_73) == 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; -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; +lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_74 = lean_ctor_get(x_73, 0); +lean_inc(x_74); +x_75 = lean_ctor_get(x_73, 1); +lean_inc(x_75); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_76 = x_73; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_73); + x_76 = lean_box(0); } -x_75 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_75, 0, x_72); -x_76 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_76, 0, x_62); -lean_ctor_set(x_76, 1, x_75); -if (lean_is_scalar(x_74)) { - x_77 = lean_alloc_ctor(0, 2, 0); +x_77 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_77, 0, x_74); +x_78 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_78, 0, x_64); +lean_ctor_set(x_78, 1, x_77); +if (lean_is_scalar(x_76)) { + x_79 = lean_alloc_ctor(0, 2, 0); } else { - x_77 = x_74; + x_79 = x_76; } -lean_ctor_set(x_77, 0, x_76); -lean_ctor_set(x_77, 1, x_73); -return x_77; +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_75); +return x_79; } else { -lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_78 = lean_ctor_get(x_71, 0); -lean_inc(x_78); -x_79 = lean_ctor_get(x_71, 1); -lean_inc(x_79); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_80 = x_71; +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; +x_80 = lean_ctor_get(x_73, 0); +lean_inc(x_80); +x_81 = lean_ctor_get(x_73, 1); +lean_inc(x_81); +if (lean_is_exclusive(x_73)) { + lean_ctor_release(x_73, 0); + lean_ctor_release(x_73, 1); + x_82 = x_73; } else { - lean_dec_ref(x_71); - x_80 = lean_box(0); + lean_dec_ref(x_73); + x_82 = lean_box(0); } -if (lean_is_scalar(x_80)) { - x_81 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_82)) { + x_83 = lean_alloc_ctor(1, 2, 0); } else { - x_81 = x_80; + x_83 = x_82; } -lean_ctor_set(x_81, 0, x_78); -lean_ctor_set(x_81, 1, x_79); -return x_81; +lean_ctor_set(x_83, 0, x_80); +lean_ctor_set(x_83, 1, x_81); +return x_83; } } else { -lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -lean_dec(x_60); +lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; +lean_dec(x_62); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_82 = lean_ctor_get(x_67, 0); -lean_inc(x_82); -x_83 = lean_ctor_get(x_67, 1); -lean_inc(x_83); -if (lean_is_exclusive(x_67)) { - lean_ctor_release(x_67, 0); - lean_ctor_release(x_67, 1); - x_84 = x_67; +x_84 = lean_ctor_get(x_69, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_69, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_69)) { + lean_ctor_release(x_69, 0); + lean_ctor_release(x_69, 1); + x_86 = x_69; } else { - lean_dec_ref(x_67); - x_84 = lean_box(0); + lean_dec_ref(x_69); + x_86 = lean_box(0); } -if (lean_is_scalar(x_84)) { - x_85 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(1, 2, 0); } else { - x_85 = x_84; + x_87 = x_86; } -lean_ctor_set(x_85, 0, x_82); -lean_ctor_set(x_85, 1, x_83); -return x_85; +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_85); +return x_87; } } else { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; -lean_dec(x_55); +lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; +lean_dec(x_56); 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_86 = lean_ctor_get(x_59, 0); -lean_inc(x_86); -x_87 = lean_ctor_get(x_59, 1); -lean_inc(x_87); -if (lean_is_exclusive(x_59)) { - lean_ctor_release(x_59, 0); - lean_ctor_release(x_59, 1); - x_88 = x_59; +x_88 = lean_ctor_get(x_61, 0); +lean_inc(x_88); +x_89 = lean_ctor_get(x_61, 1); +lean_inc(x_89); +if (lean_is_exclusive(x_61)) { + lean_ctor_release(x_61, 0); + lean_ctor_release(x_61, 1); + x_90 = x_61; } else { - lean_dec_ref(x_59); - x_88 = lean_box(0); + lean_dec_ref(x_61); + x_90 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_89 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_90)) { + x_91 = lean_alloc_ctor(1, 2, 0); } else { - x_89 = x_88; + x_91 = x_90; } -lean_ctor_set(x_89, 0, x_86); -lean_ctor_set(x_89, 1, x_87); -return x_89; +lean_ctor_set(x_91, 0, x_88); +lean_ctor_set(x_91, 1, x_89); +return x_91; } } else { -lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; -lean_dec(x_55); +lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; +lean_dec(x_56); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -15229,350 +15231,351 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_90 = lean_ctor_get(x_57, 0); -lean_inc(x_90); -x_91 = lean_ctor_get(x_57, 1); -lean_inc(x_91); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_92 = x_57; +x_92 = lean_ctor_get(x_59, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_59, 1); +lean_inc(x_93); +if (lean_is_exclusive(x_59)) { + lean_ctor_release(x_59, 0); + lean_ctor_release(x_59, 1); + x_94 = x_59; } else { - lean_dec_ref(x_57); - x_92 = lean_box(0); + lean_dec_ref(x_59); + x_94 = lean_box(0); } -if (lean_is_scalar(x_92)) { - x_93 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_94)) { + x_95 = lean_alloc_ctor(1, 2, 0); } else { - x_93 = x_92; + x_95 = x_94; } -lean_ctor_set(x_93, 0, x_90); -lean_ctor_set(x_93, 1, x_91); -return x_93; +lean_ctor_set(x_95, 0, x_92); +lean_ctor_set(x_95, 1, x_93); +return x_95; } } } } else { -uint8_t x_94; -x_94 = !lean_is_exclusive(x_10); -if (x_94 == 0) +uint8_t x_96; +x_96 = !lean_is_exclusive(x_10); +if (x_96 == 0) { -lean_object* x_95; lean_object* x_96; -x_95 = lean_ctor_get(x_10, 0); +lean_object* x_97; lean_object* x_98; +x_97 = lean_ctor_get(x_10, 0); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_96 = l_Lean_Elab_Term_elabType(x_95, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_96) == 0) +x_98 = l_Lean_Elab_Term_elabType(x_97, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_98) == 0) { -lean_object* x_97; lean_object* x_98; uint8_t x_99; lean_object* x_100; -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 = 0; +lean_object* x_99; lean_object* x_100; uint8_t x_101; uint8_t x_102; lean_object* x_103; +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 = 1; +x_102 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_100 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_99, x_99, x_3, x_4, x_5, x_6, x_7, x_8, x_98); -if (lean_obj_tag(x_100) == 0) +x_103 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_101, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_100); +if (lean_obj_tag(x_103) == 0) { -lean_object* x_101; lean_object* x_102; -x_101 = lean_ctor_get(x_100, 1); -lean_inc(x_101); -lean_dec(x_100); +lean_object* x_104; lean_object* x_105; +x_104 = lean_ctor_get(x_103, 1); +lean_inc(x_104); +lean_dec(x_103); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -x_102 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_101); -if (lean_obj_tag(x_102) == 0) +x_105 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_104); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_103; -x_103 = lean_ctor_get(x_1, 7); -lean_inc(x_103); +lean_object* x_106; +x_106 = lean_ctor_get(x_1, 7); +lean_inc(x_106); lean_dec(x_1); -if (lean_obj_tag(x_103) == 0) +if (lean_obj_tag(x_106) == 0) { -lean_object* x_104; lean_object* x_105; uint8_t x_106; uint8_t x_107; lean_object* x_108; +lean_object* x_107; lean_object* x_108; uint8_t x_109; uint8_t x_110; lean_object* x_111; lean_dec(x_4); lean_dec(x_3); -x_104 = lean_ctor_get(x_102, 0); -lean_inc(x_104); -x_105 = lean_ctor_get(x_102, 1); -lean_inc(x_105); -lean_dec(x_102); -x_106 = 1; -x_107 = 1; -x_108 = l_Lean_Meta_mkForallFVars(x_104, x_97, x_99, x_106, x_107, x_5, x_6, x_7, x_8, x_105); +x_107 = lean_ctor_get(x_105, 0); +lean_inc(x_107); +x_108 = lean_ctor_get(x_105, 1); +lean_inc(x_108); +lean_dec(x_105); +x_109 = 1; +x_110 = 1; +x_111 = l_Lean_Meta_mkForallFVars(x_107, x_99, x_102, x_109, x_110, x_5, x_6, x_7, x_8, x_108); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_108) == 0) +if (lean_obj_tag(x_111) == 0) { -uint8_t x_109; -x_109 = !lean_is_exclusive(x_108); -if (x_109 == 0) +uint8_t x_112; +x_112 = !lean_is_exclusive(x_111); +if (x_112 == 0) { -lean_object* x_110; lean_object* x_111; lean_object* x_112; -x_110 = lean_ctor_get(x_108, 0); -lean_ctor_set(x_10, 0, x_110); -x_111 = lean_box(0); -x_112 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_112, 0, x_10); -lean_ctor_set(x_112, 1, x_111); -lean_ctor_set(x_108, 0, x_112); -return x_108; +lean_object* x_113; lean_object* x_114; lean_object* x_115; +x_113 = lean_ctor_get(x_111, 0); +lean_ctor_set(x_10, 0, x_113); +x_114 = lean_box(0); +x_115 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_115, 0, x_10); +lean_ctor_set(x_115, 1, x_114); +lean_ctor_set(x_111, 0, x_115); +return x_111; } else { -lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; -x_113 = lean_ctor_get(x_108, 0); -x_114 = lean_ctor_get(x_108, 1); -lean_inc(x_114); -lean_inc(x_113); -lean_dec(x_108); -lean_ctor_set(x_10, 0, x_113); -x_115 = lean_box(0); -x_116 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_116, 0, x_10); -lean_ctor_set(x_116, 1, x_115); -x_117 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_117, 0, x_116); -lean_ctor_set(x_117, 1, x_114); -return x_117; +lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; +x_116 = lean_ctor_get(x_111, 0); +x_117 = lean_ctor_get(x_111, 1); +lean_inc(x_117); +lean_inc(x_116); +lean_dec(x_111); +lean_ctor_set(x_10, 0, x_116); +x_118 = lean_box(0); +x_119 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_119, 0, x_10); +lean_ctor_set(x_119, 1, x_118); +x_120 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_120, 0, x_119); +lean_ctor_set(x_120, 1, x_117); +return x_120; } } else { -uint8_t x_118; +uint8_t x_121; lean_free_object(x_10); -x_118 = !lean_is_exclusive(x_108); -if (x_118 == 0) +x_121 = !lean_is_exclusive(x_111); +if (x_121 == 0) { -return x_108; +return x_111; } else { -lean_object* x_119; lean_object* x_120; lean_object* x_121; -x_119 = lean_ctor_get(x_108, 0); -x_120 = lean_ctor_get(x_108, 1); -lean_inc(x_120); -lean_inc(x_119); -lean_dec(x_108); -x_121 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_121, 0, x_119); -lean_ctor_set(x_121, 1, x_120); -return x_121; +lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_122 = lean_ctor_get(x_111, 0); +x_123 = lean_ctor_get(x_111, 1); +lean_inc(x_123); +lean_inc(x_122); +lean_dec(x_111); +x_124 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_124, 0, x_122); +lean_ctor_set(x_124, 1, x_123); +return x_124; } } } else { -lean_object* x_122; lean_object* x_123; uint8_t x_124; -x_122 = lean_ctor_get(x_102, 0); -lean_inc(x_122); -x_123 = lean_ctor_get(x_102, 1); -lean_inc(x_123); -lean_dec(x_102); -x_124 = !lean_is_exclusive(x_103); -if (x_124 == 0) +lean_object* x_125; lean_object* x_126; uint8_t x_127; +x_125 = lean_ctor_get(x_105, 0); +lean_inc(x_125); +x_126 = lean_ctor_get(x_105, 1); +lean_inc(x_126); +lean_dec(x_105); +x_127 = !lean_is_exclusive(x_106); +if (x_127 == 0) { -lean_object* x_125; lean_object* x_126; uint8_t x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; -x_125 = lean_ctor_get(x_103, 0); -lean_inc(x_97); -lean_ctor_set(x_103, 0, x_97); -x_126 = lean_box(0); -x_127 = 1; -x_128 = lean_box(x_127); -x_129 = lean_box(x_127); -x_130 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_130, 0, x_125); -lean_closure_set(x_130, 1, x_103); -lean_closure_set(x_130, 2, x_128); -lean_closure_set(x_130, 3, x_129); -lean_closure_set(x_130, 4, x_126); +lean_object* x_128; lean_object* x_129; uint8_t x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; +x_128 = lean_ctor_get(x_106, 0); +lean_inc(x_99); +lean_ctor_set(x_106, 0, x_99); +x_129 = lean_box(0); +x_130 = 1; +x_131 = lean_box(x_130); +x_132 = lean_box(x_130); +x_133 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_133, 0, x_128); +lean_closure_set(x_133, 1, x_106); +lean_closure_set(x_133, 2, x_131); +lean_closure_set(x_133, 3, x_132); +lean_closure_set(x_133, 4, x_129); 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_131 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_130, x_3, x_4, x_5, x_6, x_7, x_8, x_123); -if (lean_obj_tag(x_131) == 0) +x_134 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_133, x_3, x_4, x_5, x_6, x_7, x_8, x_126); +if (lean_obj_tag(x_134) == 0) { -lean_object* x_132; lean_object* x_133; lean_object* x_134; -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); +lean_object* x_135; lean_object* x_136; lean_object* x_137; +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_134 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_99, x_99, x_3, x_4, x_5, x_6, x_7, x_8, x_133); -if (lean_obj_tag(x_134) == 0) -{ -lean_object* x_135; uint8_t x_136; lean_object* x_137; -x_135 = lean_ctor_get(x_134, 1); -lean_inc(x_135); -lean_dec(x_134); -x_136 = 1; -lean_inc(x_122); -x_137 = l_Lean_Meta_mkForallFVars(x_122, x_97, x_99, x_127, x_136, x_5, x_6, x_7, x_8, x_135); +x_137 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_101, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_136); if (lean_obj_tag(x_137) == 0) { -lean_object* x_138; lean_object* x_139; lean_object* x_140; -x_138 = lean_ctor_get(x_137, 0); +lean_object* x_138; uint8_t x_139; lean_object* x_140; +x_138 = lean_ctor_get(x_137, 1); lean_inc(x_138); -x_139 = lean_ctor_get(x_137, 1); -lean_inc(x_139); lean_dec(x_137); -x_140 = l_Lean_Meta_mkLambdaFVars(x_122, x_132, x_99, x_127, x_136, x_5, x_6, x_7, x_8, x_139); +x_139 = 1; +lean_inc(x_125); +x_140 = l_Lean_Meta_mkForallFVars(x_125, x_99, x_102, x_130, x_139, x_5, x_6, x_7, x_8, x_138); +if (lean_obj_tag(x_140) == 0) +{ +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_140, 0); +lean_inc(x_141); +x_142 = lean_ctor_get(x_140, 1); +lean_inc(x_142); +lean_dec(x_140); +x_143 = l_Lean_Meta_mkLambdaFVars(x_125, x_135, x_102, x_130, x_139, x_5, x_6, x_7, x_8, x_142); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_140) == 0) +if (lean_obj_tag(x_143) == 0) { -uint8_t x_141; -x_141 = !lean_is_exclusive(x_140); -if (x_141 == 0) -{ -lean_object* x_142; lean_object* x_143; lean_object* x_144; -x_142 = lean_ctor_get(x_140, 0); -lean_ctor_set(x_10, 0, x_138); -x_143 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_143, 0, x_142); -x_144 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_144, 0, x_10); -lean_ctor_set(x_144, 1, x_143); -lean_ctor_set(x_140, 0, x_144); -return x_140; +uint8_t x_144; +x_144 = !lean_is_exclusive(x_143); +if (x_144 == 0) +{ +lean_object* x_145; lean_object* x_146; lean_object* x_147; +x_145 = lean_ctor_get(x_143, 0); +lean_ctor_set(x_10, 0, x_141); +x_146 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_146, 0, x_145); +x_147 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_147, 0, x_10); +lean_ctor_set(x_147, 1, x_146); +lean_ctor_set(x_143, 0, x_147); +return x_143; } else { -lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; -x_145 = lean_ctor_get(x_140, 0); -x_146 = lean_ctor_get(x_140, 1); -lean_inc(x_146); -lean_inc(x_145); -lean_dec(x_140); -lean_ctor_set(x_10, 0, x_138); -x_147 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_147, 0, x_145); -x_148 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_148, 0, x_10); -lean_ctor_set(x_148, 1, x_147); -x_149 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_149, 0, x_148); -lean_ctor_set(x_149, 1, x_146); -return x_149; +lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; +x_148 = lean_ctor_get(x_143, 0); +x_149 = lean_ctor_get(x_143, 1); +lean_inc(x_149); +lean_inc(x_148); +lean_dec(x_143); +lean_ctor_set(x_10, 0, x_141); +x_150 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_150, 0, x_148); +x_151 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_151, 0, x_10); +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_149); +return x_152; } } else { -uint8_t x_150; -lean_dec(x_138); +uint8_t x_153; +lean_dec(x_141); lean_free_object(x_10); -x_150 = !lean_is_exclusive(x_140); -if (x_150 == 0) +x_153 = !lean_is_exclusive(x_143); +if (x_153 == 0) { -return x_140; +return x_143; } else { -lean_object* x_151; lean_object* x_152; lean_object* x_153; -x_151 = lean_ctor_get(x_140, 0); -x_152 = lean_ctor_get(x_140, 1); -lean_inc(x_152); -lean_inc(x_151); -lean_dec(x_140); -x_153 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_153, 0, x_151); -lean_ctor_set(x_153, 1, x_152); -return x_153; +lean_object* x_154; lean_object* x_155; lean_object* x_156; +x_154 = lean_ctor_get(x_143, 0); +x_155 = lean_ctor_get(x_143, 1); +lean_inc(x_155); +lean_inc(x_154); +lean_dec(x_143); +x_156 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_156, 0, x_154); +lean_ctor_set(x_156, 1, x_155); +return x_156; } } } else { -uint8_t x_154; -lean_dec(x_132); -lean_dec(x_122); +uint8_t x_157; +lean_dec(x_135); +lean_dec(x_125); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_154 = !lean_is_exclusive(x_137); -if (x_154 == 0) +x_157 = !lean_is_exclusive(x_140); +if (x_157 == 0) { -return x_137; +return x_140; } else { -lean_object* x_155; lean_object* x_156; lean_object* x_157; -x_155 = lean_ctor_get(x_137, 0); -x_156 = lean_ctor_get(x_137, 1); -lean_inc(x_156); -lean_inc(x_155); -lean_dec(x_137); -x_157 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_157, 0, x_155); -lean_ctor_set(x_157, 1, x_156); -return x_157; +lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_158 = lean_ctor_get(x_140, 0); +x_159 = lean_ctor_get(x_140, 1); +lean_inc(x_159); +lean_inc(x_158); +lean_dec(x_140); +x_160 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_160, 0, x_158); +lean_ctor_set(x_160, 1, x_159); +return x_160; } } } else { -uint8_t x_158; -lean_dec(x_132); -lean_dec(x_122); -lean_dec(x_97); +uint8_t x_161; +lean_dec(x_135); +lean_dec(x_125); +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_158 = !lean_is_exclusive(x_134); -if (x_158 == 0) +x_161 = !lean_is_exclusive(x_137); +if (x_161 == 0) { -return x_134; +return x_137; } else { -lean_object* x_159; lean_object* x_160; lean_object* x_161; -x_159 = lean_ctor_get(x_134, 0); -x_160 = lean_ctor_get(x_134, 1); -lean_inc(x_160); -lean_inc(x_159); -lean_dec(x_134); -x_161 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_161, 0, x_159); -lean_ctor_set(x_161, 1, x_160); -return x_161; +lean_object* x_162; lean_object* x_163; lean_object* x_164; +x_162 = lean_ctor_get(x_137, 0); +x_163 = lean_ctor_get(x_137, 1); +lean_inc(x_163); +lean_inc(x_162); +lean_dec(x_137); +x_164 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_164, 0, x_162); +lean_ctor_set(x_164, 1, x_163); +return x_164; } } } else { -uint8_t x_162; -lean_dec(x_122); -lean_dec(x_97); +uint8_t x_165; +lean_dec(x_125); +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); @@ -15580,214 +15583,214 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_162 = !lean_is_exclusive(x_131); -if (x_162 == 0) +x_165 = !lean_is_exclusive(x_134); +if (x_165 == 0) { -return x_131; +return x_134; } else { -lean_object* x_163; lean_object* x_164; lean_object* x_165; -x_163 = lean_ctor_get(x_131, 0); -x_164 = lean_ctor_get(x_131, 1); -lean_inc(x_164); -lean_inc(x_163); -lean_dec(x_131); -x_165 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_165, 0, x_163); -lean_ctor_set(x_165, 1, x_164); -return x_165; +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_134, 0); +x_167 = lean_ctor_get(x_134, 1); +lean_inc(x_167); +lean_inc(x_166); +lean_dec(x_134); +x_168 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_168, 0, x_166); +lean_ctor_set(x_168, 1, x_167); +return x_168; } } } else { -lean_object* x_166; lean_object* x_167; lean_object* x_168; uint8_t x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; -x_166 = lean_ctor_get(x_103, 0); -lean_inc(x_166); -lean_dec(x_103); -lean_inc(x_97); -x_167 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_167, 0, x_97); -x_168 = lean_box(0); -x_169 = 1; -x_170 = lean_box(x_169); -x_171 = lean_box(x_169); -x_172 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_172, 0, x_166); -lean_closure_set(x_172, 1, x_167); -lean_closure_set(x_172, 2, x_170); -lean_closure_set(x_172, 3, x_171); -lean_closure_set(x_172, 4, x_168); +lean_object* x_169; lean_object* x_170; lean_object* x_171; uint8_t x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; +x_169 = lean_ctor_get(x_106, 0); +lean_inc(x_169); +lean_dec(x_106); +lean_inc(x_99); +x_170 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_170, 0, x_99); +x_171 = lean_box(0); +x_172 = 1; +x_173 = lean_box(x_172); +x_174 = lean_box(x_172); +x_175 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_175, 0, x_169); +lean_closure_set(x_175, 1, x_170); +lean_closure_set(x_175, 2, x_173); +lean_closure_set(x_175, 3, x_174); +lean_closure_set(x_175, 4, x_171); 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_173 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_172, x_3, x_4, x_5, x_6, x_7, x_8, x_123); -if (lean_obj_tag(x_173) == 0) +x_176 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_175, x_3, x_4, x_5, x_6, x_7, x_8, x_126); +if (lean_obj_tag(x_176) == 0) { -lean_object* x_174; lean_object* x_175; lean_object* x_176; -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); +lean_object* x_177; lean_object* x_178; lean_object* x_179; +x_177 = lean_ctor_get(x_176, 0); +lean_inc(x_177); +x_178 = lean_ctor_get(x_176, 1); +lean_inc(x_178); +lean_dec(x_176); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_176 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_99, x_99, x_3, x_4, x_5, x_6, x_7, x_8, x_175); -if (lean_obj_tag(x_176) == 0) -{ -lean_object* x_177; uint8_t x_178; lean_object* x_179; -x_177 = lean_ctor_get(x_176, 1); -lean_inc(x_177); -lean_dec(x_176); -x_178 = 1; -lean_inc(x_122); -x_179 = l_Lean_Meta_mkForallFVars(x_122, x_97, x_99, x_169, x_178, x_5, x_6, x_7, x_8, x_177); +x_179 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_101, x_102, x_3, x_4, x_5, x_6, x_7, x_8, x_178); if (lean_obj_tag(x_179) == 0) { -lean_object* x_180; lean_object* x_181; lean_object* x_182; -x_180 = lean_ctor_get(x_179, 0); +lean_object* x_180; uint8_t x_181; lean_object* x_182; +x_180 = lean_ctor_get(x_179, 1); lean_inc(x_180); -x_181 = lean_ctor_get(x_179, 1); -lean_inc(x_181); lean_dec(x_179); -x_182 = l_Lean_Meta_mkLambdaFVars(x_122, x_174, x_99, x_169, x_178, x_5, x_6, x_7, x_8, x_181); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); +x_181 = 1; +lean_inc(x_125); +x_182 = l_Lean_Meta_mkForallFVars(x_125, x_99, x_102, x_172, x_181, x_5, x_6, x_7, x_8, x_180); if (lean_obj_tag(x_182) == 0) { -lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +lean_object* x_183; lean_object* x_184; lean_object* x_185; x_183 = lean_ctor_get(x_182, 0); lean_inc(x_183); x_184 = lean_ctor_get(x_182, 1); lean_inc(x_184); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_185 = x_182; +lean_dec(x_182); +x_185 = l_Lean_Meta_mkLambdaFVars(x_125, x_177, x_102, x_172, x_181, x_5, x_6, x_7, x_8, x_184); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +if (lean_obj_tag(x_185) == 0) +{ +lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; +x_186 = lean_ctor_get(x_185, 0); +lean_inc(x_186); +x_187 = lean_ctor_get(x_185, 1); +lean_inc(x_187); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_188 = x_185; } else { - lean_dec_ref(x_182); - x_185 = lean_box(0); -} -lean_ctor_set(x_10, 0, x_180); -x_186 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_186, 0, x_183); -x_187 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_187, 0, x_10); -lean_ctor_set(x_187, 1, x_186); -if (lean_is_scalar(x_185)) { - x_188 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_185); + x_188 = lean_box(0); +} +lean_ctor_set(x_10, 0, x_183); +x_189 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_189, 0, x_186); +x_190 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_190, 0, x_10); +lean_ctor_set(x_190, 1, x_189); +if (lean_is_scalar(x_188)) { + x_191 = lean_alloc_ctor(0, 2, 0); } else { - x_188 = x_185; + x_191 = x_188; } -lean_ctor_set(x_188, 0, x_187); -lean_ctor_set(x_188, 1, x_184); -return x_188; +lean_ctor_set(x_191, 0, x_190); +lean_ctor_set(x_191, 1, x_187); +return x_191; } else { -lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; -lean_dec(x_180); +lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; +lean_dec(x_183); lean_free_object(x_10); -x_189 = lean_ctor_get(x_182, 0); -lean_inc(x_189); -x_190 = lean_ctor_get(x_182, 1); -lean_inc(x_190); -if (lean_is_exclusive(x_182)) { - lean_ctor_release(x_182, 0); - lean_ctor_release(x_182, 1); - x_191 = x_182; +x_192 = lean_ctor_get(x_185, 0); +lean_inc(x_192); +x_193 = lean_ctor_get(x_185, 1); +lean_inc(x_193); +if (lean_is_exclusive(x_185)) { + lean_ctor_release(x_185, 0); + lean_ctor_release(x_185, 1); + x_194 = x_185; } else { - lean_dec_ref(x_182); - x_191 = lean_box(0); + lean_dec_ref(x_185); + x_194 = lean_box(0); } -if (lean_is_scalar(x_191)) { - x_192 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_194)) { + x_195 = lean_alloc_ctor(1, 2, 0); } else { - x_192 = x_191; + x_195 = x_194; } -lean_ctor_set(x_192, 0, x_189); -lean_ctor_set(x_192, 1, x_190); -return x_192; +lean_ctor_set(x_195, 0, x_192); +lean_ctor_set(x_195, 1, x_193); +return x_195; } } else { -lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; -lean_dec(x_174); -lean_dec(x_122); +lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; +lean_dec(x_177); +lean_dec(x_125); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_193 = lean_ctor_get(x_179, 0); -lean_inc(x_193); -x_194 = lean_ctor_get(x_179, 1); -lean_inc(x_194); -if (lean_is_exclusive(x_179)) { - lean_ctor_release(x_179, 0); - lean_ctor_release(x_179, 1); - x_195 = x_179; +x_196 = lean_ctor_get(x_182, 0); +lean_inc(x_196); +x_197 = lean_ctor_get(x_182, 1); +lean_inc(x_197); +if (lean_is_exclusive(x_182)) { + lean_ctor_release(x_182, 0); + lean_ctor_release(x_182, 1); + x_198 = x_182; } else { - lean_dec_ref(x_179); - x_195 = lean_box(0); + lean_dec_ref(x_182); + x_198 = lean_box(0); } -if (lean_is_scalar(x_195)) { - x_196 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_198)) { + x_199 = lean_alloc_ctor(1, 2, 0); } else { - x_196 = x_195; + x_199 = x_198; } -lean_ctor_set(x_196, 0, x_193); -lean_ctor_set(x_196, 1, x_194); -return x_196; +lean_ctor_set(x_199, 0, x_196); +lean_ctor_set(x_199, 1, x_197); +return x_199; } } else { -lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; -lean_dec(x_174); -lean_dec(x_122); -lean_dec(x_97); +lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; +lean_dec(x_177); +lean_dec(x_125); +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_197 = lean_ctor_get(x_176, 0); -lean_inc(x_197); -x_198 = lean_ctor_get(x_176, 1); -lean_inc(x_198); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_199 = x_176; +x_200 = lean_ctor_get(x_179, 0); +lean_inc(x_200); +x_201 = lean_ctor_get(x_179, 1); +lean_inc(x_201); +if (lean_is_exclusive(x_179)) { + lean_ctor_release(x_179, 0); + lean_ctor_release(x_179, 1); + x_202 = x_179; } else { - lean_dec_ref(x_176); - x_199 = lean_box(0); + lean_dec_ref(x_179); + x_202 = lean_box(0); } -if (lean_is_scalar(x_199)) { - x_200 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_202)) { + x_203 = lean_alloc_ctor(1, 2, 0); } else { - x_200 = x_199; + x_203 = x_202; } -lean_ctor_set(x_200, 0, x_197); -lean_ctor_set(x_200, 1, x_198); -return x_200; +lean_ctor_set(x_203, 0, x_200); +lean_ctor_set(x_203, 1, x_201); +return x_203; } } else { -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; -lean_dec(x_122); -lean_dec(x_97); +lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; +lean_dec(x_125); +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); @@ -15795,34 +15798,34 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_201 = lean_ctor_get(x_173, 0); -lean_inc(x_201); -x_202 = lean_ctor_get(x_173, 1); -lean_inc(x_202); -if (lean_is_exclusive(x_173)) { - lean_ctor_release(x_173, 0); - lean_ctor_release(x_173, 1); - x_203 = x_173; +x_204 = lean_ctor_get(x_176, 0); +lean_inc(x_204); +x_205 = lean_ctor_get(x_176, 1); +lean_inc(x_205); +if (lean_is_exclusive(x_176)) { + lean_ctor_release(x_176, 0); + lean_ctor_release(x_176, 1); + x_206 = x_176; } else { - lean_dec_ref(x_173); - x_203 = lean_box(0); + lean_dec_ref(x_176); + x_206 = lean_box(0); } -if (lean_is_scalar(x_203)) { - x_204 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_206)) { + x_207 = lean_alloc_ctor(1, 2, 0); } else { - x_204 = x_203; + x_207 = x_206; } -lean_ctor_set(x_204, 0, x_201); -lean_ctor_set(x_204, 1, x_202); -return x_204; +lean_ctor_set(x_207, 0, x_204); +lean_ctor_set(x_207, 1, x_205); +return x_207; } } } } else { -uint8_t x_205; -lean_dec(x_97); +uint8_t x_208; +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); @@ -15831,30 +15834,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_205 = !lean_is_exclusive(x_102); -if (x_205 == 0) +x_208 = !lean_is_exclusive(x_105); +if (x_208 == 0) { -return x_102; +return x_105; } else { -lean_object* x_206; lean_object* x_207; lean_object* x_208; -x_206 = lean_ctor_get(x_102, 0); -x_207 = lean_ctor_get(x_102, 1); -lean_inc(x_207); -lean_inc(x_206); -lean_dec(x_102); -x_208 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); -return x_208; +lean_object* x_209; lean_object* x_210; lean_object* x_211; +x_209 = lean_ctor_get(x_105, 0); +x_210 = lean_ctor_get(x_105, 1); +lean_inc(x_210); +lean_inc(x_209); +lean_dec(x_105); +x_211 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_211, 0, x_209); +lean_ctor_set(x_211, 1, x_210); +return x_211; } } } else { -uint8_t x_209; -lean_dec(x_97); +uint8_t x_212; +lean_dec(x_99); lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); @@ -15864,29 +15867,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_209 = !lean_is_exclusive(x_100); -if (x_209 == 0) +x_212 = !lean_is_exclusive(x_103); +if (x_212 == 0) { -return x_100; +return x_103; } else { -lean_object* x_210; lean_object* x_211; lean_object* x_212; -x_210 = lean_ctor_get(x_100, 0); -x_211 = lean_ctor_get(x_100, 1); -lean_inc(x_211); -lean_inc(x_210); -lean_dec(x_100); -x_212 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_212, 0, x_210); -lean_ctor_set(x_212, 1, x_211); -return x_212; +lean_object* x_213; lean_object* x_214; lean_object* x_215; +x_213 = lean_ctor_get(x_103, 0); +x_214 = lean_ctor_get(x_103, 1); +lean_inc(x_214); +lean_inc(x_213); +lean_dec(x_103); +x_215 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_215, 0, x_213); +lean_ctor_set(x_215, 1, x_214); +return x_215; } } } else { -uint8_t x_213; +uint8_t x_216; lean_free_object(x_10); lean_dec(x_8); lean_dec(x_7); @@ -15896,31 +15899,31 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_213 = !lean_is_exclusive(x_96); -if (x_213 == 0) +x_216 = !lean_is_exclusive(x_98); +if (x_216 == 0) { -return x_96; +return x_98; } else { -lean_object* x_214; lean_object* x_215; lean_object* x_216; -x_214 = lean_ctor_get(x_96, 0); -x_215 = lean_ctor_get(x_96, 1); -lean_inc(x_215); -lean_inc(x_214); -lean_dec(x_96); -x_216 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_216, 0, x_214); -lean_ctor_set(x_216, 1, x_215); -return x_216; +lean_object* x_217; lean_object* x_218; lean_object* x_219; +x_217 = lean_ctor_get(x_98, 0); +x_218 = lean_ctor_get(x_98, 1); +lean_inc(x_218); +lean_inc(x_217); +lean_dec(x_98); +x_219 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_219, 0, x_217); +lean_ctor_set(x_219, 1, x_218); +return x_219; } } } else { -lean_object* x_217; lean_object* x_218; -x_217 = lean_ctor_get(x_10, 0); -lean_inc(x_217); +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_10, 0); +lean_inc(x_220); lean_dec(x_10); lean_inc(x_8); lean_inc(x_7); @@ -15928,265 +15931,235 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_218 = l_Lean_Elab_Term_elabType(x_217, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_218) == 0) +x_221 = l_Lean_Elab_Term_elabType(x_220, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_221) == 0) { -lean_object* x_219; lean_object* x_220; uint8_t x_221; lean_object* x_222; -x_219 = lean_ctor_get(x_218, 0); -lean_inc(x_219); -x_220 = lean_ctor_get(x_218, 1); -lean_inc(x_220); -lean_dec(x_218); -x_221 = 0; +lean_object* x_222; lean_object* x_223; uint8_t x_224; uint8_t x_225; lean_object* x_226; +x_222 = lean_ctor_get(x_221, 0); +lean_inc(x_222); +x_223 = lean_ctor_get(x_221, 1); +lean_inc(x_223); +lean_dec(x_221); +x_224 = 1; +x_225 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_222 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_221, x_221, x_3, x_4, x_5, x_6, x_7, x_8, x_220); -if (lean_obj_tag(x_222) == 0) +x_226 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_224, x_225, x_3, x_4, x_5, x_6, x_7, x_8, x_223); +if (lean_obj_tag(x_226) == 0) { -lean_object* x_223; lean_object* x_224; -x_223 = lean_ctor_get(x_222, 1); -lean_inc(x_223); -lean_dec(x_222); +lean_object* x_227; lean_object* x_228; +x_227 = lean_ctor_get(x_226, 1); +lean_inc(x_227); +lean_dec(x_226); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_3); -x_224 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_223); -if (lean_obj_tag(x_224) == 0) +x_228 = l_Lean_Elab_Term_addAutoBoundImplicits(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_227); +if (lean_obj_tag(x_228) == 0) { -lean_object* x_225; -x_225 = lean_ctor_get(x_1, 7); -lean_inc(x_225); +lean_object* x_229; +x_229 = lean_ctor_get(x_1, 7); +lean_inc(x_229); lean_dec(x_1); -if (lean_obj_tag(x_225) == 0) +if (lean_obj_tag(x_229) == 0) { -lean_object* x_226; lean_object* x_227; uint8_t x_228; uint8_t x_229; lean_object* x_230; +lean_object* x_230; lean_object* x_231; uint8_t x_232; uint8_t x_233; lean_object* x_234; lean_dec(x_4); lean_dec(x_3); -x_226 = lean_ctor_get(x_224, 0); -lean_inc(x_226); -x_227 = lean_ctor_get(x_224, 1); -lean_inc(x_227); -lean_dec(x_224); -x_228 = 1; -x_229 = 1; -x_230 = l_Lean_Meta_mkForallFVars(x_226, x_219, x_221, x_228, x_229, x_5, x_6, x_7, x_8, x_227); +x_230 = lean_ctor_get(x_228, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_228, 1); +lean_inc(x_231); +lean_dec(x_228); +x_232 = 1; +x_233 = 1; +x_234 = l_Lean_Meta_mkForallFVars(x_230, x_222, x_225, x_232, x_233, x_5, x_6, x_7, x_8, x_231); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_230) == 0) +if (lean_obj_tag(x_234) == 0) { -lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_ctor_get(x_230, 1); -lean_inc(x_232); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_233 = x_230; +lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; +x_235 = lean_ctor_get(x_234, 0); +lean_inc(x_235); +x_236 = lean_ctor_get(x_234, 1); +lean_inc(x_236); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + x_237 = x_234; } else { - lean_dec_ref(x_230); - x_233 = lean_box(0); -} -x_234 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_234, 0, x_231); -x_235 = lean_box(0); -x_236 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_236, 0, x_234); -lean_ctor_set(x_236, 1, x_235); -if (lean_is_scalar(x_233)) { - x_237 = lean_alloc_ctor(0, 2, 0); + lean_dec_ref(x_234); + x_237 = lean_box(0); +} +x_238 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_238, 0, x_235); +x_239 = lean_box(0); +x_240 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_240, 0, x_238); +lean_ctor_set(x_240, 1, x_239); +if (lean_is_scalar(x_237)) { + x_241 = lean_alloc_ctor(0, 2, 0); } else { - x_237 = x_233; + x_241 = x_237; } -lean_ctor_set(x_237, 0, x_236); -lean_ctor_set(x_237, 1, x_232); -return x_237; +lean_ctor_set(x_241, 0, x_240); +lean_ctor_set(x_241, 1, x_236); +return x_241; } else { -lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; -x_238 = lean_ctor_get(x_230, 0); -lean_inc(x_238); -x_239 = lean_ctor_get(x_230, 1); -lean_inc(x_239); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_240 = x_230; +lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_242 = lean_ctor_get(x_234, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_234, 1); +lean_inc(x_243); +if (lean_is_exclusive(x_234)) { + lean_ctor_release(x_234, 0); + lean_ctor_release(x_234, 1); + x_244 = x_234; } else { - lean_dec_ref(x_230); - x_240 = lean_box(0); + lean_dec_ref(x_234); + x_244 = lean_box(0); } -if (lean_is_scalar(x_240)) { - x_241 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_244)) { + x_245 = lean_alloc_ctor(1, 2, 0); } else { - x_241 = x_240; + x_245 = x_244; } -lean_ctor_set(x_241, 0, x_238); -lean_ctor_set(x_241, 1, x_239); -return x_241; +lean_ctor_set(x_245, 0, x_242); +lean_ctor_set(x_245, 1, x_243); +return x_245; } } else { -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; uint8_t x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; -x_242 = lean_ctor_get(x_224, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_224, 1); -lean_inc(x_243); -lean_dec(x_224); -x_244 = lean_ctor_get(x_225, 0); -lean_inc(x_244); -if (lean_is_exclusive(x_225)) { - lean_ctor_release(x_225, 0); - x_245 = x_225; +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; uint8_t x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_246 = lean_ctor_get(x_228, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_228, 1); +lean_inc(x_247); +lean_dec(x_228); +x_248 = lean_ctor_get(x_229, 0); +lean_inc(x_248); +if (lean_is_exclusive(x_229)) { + lean_ctor_release(x_229, 0); + x_249 = x_229; } else { - lean_dec_ref(x_225); - x_245 = lean_box(0); + lean_dec_ref(x_229); + x_249 = lean_box(0); } -lean_inc(x_219); -if (lean_is_scalar(x_245)) { - x_246 = lean_alloc_ctor(1, 1, 0); +lean_inc(x_222); +if (lean_is_scalar(x_249)) { + x_250 = lean_alloc_ctor(1, 1, 0); } else { - x_246 = x_245; -} -lean_ctor_set(x_246, 0, x_219); -x_247 = lean_box(0); -x_248 = 1; -x_249 = lean_box(x_248); -x_250 = lean_box(x_248); -x_251 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); -lean_closure_set(x_251, 0, x_244); -lean_closure_set(x_251, 1, x_246); -lean_closure_set(x_251, 2, x_249); -lean_closure_set(x_251, 3, x_250); -lean_closure_set(x_251, 4, x_247); + x_250 = x_249; +} +lean_ctor_set(x_250, 0, x_222); +x_251 = lean_box(0); +x_252 = 1; +x_253 = lean_box(x_252); +x_254 = lean_box(x_252); +x_255 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabTermEnsuringType___boxed), 12, 5); +lean_closure_set(x_255, 0, x_248); +lean_closure_set(x_255, 1, x_250); +lean_closure_set(x_255, 2, x_253); +lean_closure_set(x_255, 3, x_254); +lean_closure_set(x_255, 4, x_251); 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_252 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_251, x_3, x_4, x_5, x_6, x_7, x_8, x_243); -if (lean_obj_tag(x_252) == 0) -{ -lean_object* x_253; lean_object* x_254; lean_object* x_255; -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); +x_256 = l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(x_255, x_3, x_4, x_5, x_6, x_7, x_8, x_247); +if (lean_obj_tag(x_256) == 0) +{ +lean_object* x_257; lean_object* x_258; lean_object* x_259; +x_257 = lean_ctor_get(x_256, 0); +lean_inc(x_257); +x_258 = lean_ctor_get(x_256, 1); +lean_inc(x_258); +lean_dec(x_256); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_255 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_221, x_221, x_3, x_4, x_5, x_6, x_7, x_8, x_254); -if (lean_obj_tag(x_255) == 0) -{ -lean_object* x_256; uint8_t x_257; lean_object* x_258; -x_256 = lean_ctor_get(x_255, 1); -lean_inc(x_256); -lean_dec(x_255); -x_257 = 1; -lean_inc(x_242); -x_258 = l_Lean_Meta_mkForallFVars(x_242, x_219, x_221, x_248, x_257, x_5, x_6, x_7, x_8, x_256); -if (lean_obj_tag(x_258) == 0) +x_259 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_224, x_225, x_3, x_4, x_5, x_6, x_7, x_8, x_258); +if (lean_obj_tag(x_259) == 0) { -lean_object* x_259; lean_object* x_260; lean_object* x_261; -x_259 = lean_ctor_get(x_258, 0); -lean_inc(x_259); -x_260 = lean_ctor_get(x_258, 1); +lean_object* x_260; uint8_t x_261; lean_object* x_262; +x_260 = lean_ctor_get(x_259, 1); lean_inc(x_260); -lean_dec(x_258); -x_261 = l_Lean_Meta_mkLambdaFVars(x_242, x_253, x_221, x_248, x_257, x_5, x_6, x_7, x_8, x_260); +lean_dec(x_259); +x_261 = 1; +lean_inc(x_246); +x_262 = l_Lean_Meta_mkForallFVars(x_246, x_222, x_225, x_252, x_261, x_5, x_6, x_7, x_8, x_260); +if (lean_obj_tag(x_262) == 0) +{ +lean_object* x_263; lean_object* x_264; lean_object* x_265; +x_263 = lean_ctor_get(x_262, 0); +lean_inc(x_263); +x_264 = lean_ctor_get(x_262, 1); +lean_inc(x_264); +lean_dec(x_262); +x_265 = l_Lean_Meta_mkLambdaFVars(x_246, x_257, x_225, x_252, x_261, x_5, x_6, x_7, x_8, x_264); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -if (lean_obj_tag(x_261) == 0) -{ -lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -x_262 = lean_ctor_get(x_261, 0); -lean_inc(x_262); -x_263 = lean_ctor_get(x_261, 1); -lean_inc(x_263); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_264 = x_261; -} else { - lean_dec_ref(x_261); - x_264 = lean_box(0); -} -x_265 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_265, 0, x_259); -x_266 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_266, 0, x_262); -x_267 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_267, 0, x_265); -lean_ctor_set(x_267, 1, x_266); -if (lean_is_scalar(x_264)) { - x_268 = lean_alloc_ctor(0, 2, 0); -} else { - x_268 = x_264; -} -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_263); -return x_268; -} -else +if (lean_obj_tag(x_265) == 0) { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -lean_dec(x_259); -x_269 = lean_ctor_get(x_261, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_261, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_261)) { - lean_ctor_release(x_261, 0); - lean_ctor_release(x_261, 1); - x_271 = x_261; +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; +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_265, 1); +lean_inc(x_267); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_268 = x_265; } else { - lean_dec_ref(x_261); - x_271 = lean_box(0); -} -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(1, 2, 0); + lean_dec_ref(x_265); + x_268 = lean_box(0); +} +x_269 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_269, 0, x_263); +x_270 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_270, 0, x_266); +x_271 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_271, 0, x_269); +lean_ctor_set(x_271, 1, x_270); +if (lean_is_scalar(x_268)) { + x_272 = lean_alloc_ctor(0, 2, 0); } else { - x_272 = x_271; + x_272 = x_268; } -lean_ctor_set(x_272, 0, x_269); -lean_ctor_set(x_272, 1, x_270); +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_267); return x_272; } -} else { lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; -lean_dec(x_253); -lean_dec(x_242); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -x_273 = lean_ctor_get(x_258, 0); +lean_dec(x_263); +x_273 = lean_ctor_get(x_265, 0); lean_inc(x_273); -x_274 = lean_ctor_get(x_258, 1); +x_274 = lean_ctor_get(x_265, 1); lean_inc(x_274); -if (lean_is_exclusive(x_258)) { - lean_ctor_release(x_258, 0); - lean_ctor_release(x_258, 1); - x_275 = x_258; +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_275 = x_265; } else { - lean_dec_ref(x_258); + lean_dec_ref(x_265); x_275 = lean_box(0); } if (lean_is_scalar(x_275)) { @@ -16202,23 +16175,22 @@ return x_276; else { lean_object* x_277; lean_object* x_278; lean_object* x_279; lean_object* x_280; -lean_dec(x_253); -lean_dec(x_242); -lean_dec(x_219); +lean_dec(x_257); +lean_dec(x_246); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_277 = lean_ctor_get(x_255, 0); +x_277 = lean_ctor_get(x_262, 0); lean_inc(x_277); -x_278 = lean_ctor_get(x_255, 1); +x_278 = lean_ctor_get(x_262, 1); lean_inc(x_278); -if (lean_is_exclusive(x_255)) { - lean_ctor_release(x_255, 0); - lean_ctor_release(x_255, 1); - x_279 = x_255; +if (lean_is_exclusive(x_262)) { + lean_ctor_release(x_262, 0); + lean_ctor_release(x_262, 1); + x_279 = x_262; } else { - lean_dec_ref(x_255); + lean_dec_ref(x_262); x_279 = lean_box(0); } if (lean_is_scalar(x_279)) { @@ -16234,24 +16206,23 @@ return x_280; else { lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; -lean_dec(x_242); -lean_dec(x_219); +lean_dec(x_257); +lean_dec(x_246); +lean_dec(x_222); 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_281 = lean_ctor_get(x_252, 0); +x_281 = lean_ctor_get(x_259, 0); lean_inc(x_281); -x_282 = lean_ctor_get(x_252, 1); +x_282 = lean_ctor_get(x_259, 1); lean_inc(x_282); -if (lean_is_exclusive(x_252)) { - lean_ctor_release(x_252, 0); - lean_ctor_release(x_252, 1); - x_283 = x_252; +if (lean_is_exclusive(x_259)) { + lean_ctor_release(x_259, 0); + lean_ctor_release(x_259, 1); + x_283 = x_259; } else { - lean_dec_ref(x_252); + lean_dec_ref(x_259); x_283 = lean_box(0); } if (lean_is_scalar(x_283)) { @@ -16264,28 +16235,27 @@ lean_ctor_set(x_284, 1, x_282); return x_284; } } -} else { lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; -lean_dec(x_219); +lean_dec(x_246); +lean_dec(x_222); 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_285 = lean_ctor_get(x_224, 0); +x_285 = lean_ctor_get(x_256, 0); lean_inc(x_285); -x_286 = lean_ctor_get(x_224, 1); +x_286 = lean_ctor_get(x_256, 1); lean_inc(x_286); -if (lean_is_exclusive(x_224)) { - lean_ctor_release(x_224, 0); - lean_ctor_release(x_224, 1); - x_287 = x_224; +if (lean_is_exclusive(x_256)) { + lean_ctor_release(x_256, 0); + lean_ctor_release(x_256, 1); + x_287 = x_256; } else { - lean_dec_ref(x_224); + lean_dec_ref(x_256); x_287 = lean_box(0); } if (lean_is_scalar(x_287)) { @@ -16298,28 +16268,28 @@ lean_ctor_set(x_288, 1, x_286); return x_288; } } +} else { lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; -lean_dec(x_219); +lean_dec(x_222); 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_289 = lean_ctor_get(x_222, 0); +x_289 = lean_ctor_get(x_228, 0); lean_inc(x_289); -x_290 = lean_ctor_get(x_222, 1); +x_290 = lean_ctor_get(x_228, 1); lean_inc(x_290); -if (lean_is_exclusive(x_222)) { - lean_ctor_release(x_222, 0); - lean_ctor_release(x_222, 1); - x_291 = x_222; +if (lean_is_exclusive(x_228)) { + lean_ctor_release(x_228, 0); + lean_ctor_release(x_228, 1); + x_291 = x_228; } else { - lean_dec_ref(x_222); + lean_dec_ref(x_228); x_291 = lean_box(0); } if (lean_is_scalar(x_291)) { @@ -16335,6 +16305,7 @@ return x_292; else { lean_object* x_293; lean_object* x_294; lean_object* x_295; lean_object* x_296; +lean_dec(x_222); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -16343,16 +16314,16 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_293 = lean_ctor_get(x_218, 0); +x_293 = lean_ctor_get(x_226, 0); lean_inc(x_293); -x_294 = lean_ctor_get(x_218, 1); +x_294 = lean_ctor_get(x_226, 1); lean_inc(x_294); -if (lean_is_exclusive(x_218)) { - lean_ctor_release(x_218, 0); - lean_ctor_release(x_218, 1); - x_295 = x_218; +if (lean_is_exclusive(x_226)) { + lean_ctor_release(x_226, 0); + lean_ctor_release(x_226, 1); + x_295 = x_226; } else { - lean_dec_ref(x_218); + lean_dec_ref(x_226); x_295 = lean_box(0); } if (lean_is_scalar(x_295)) { @@ -16365,6 +16336,39 @@ lean_ctor_set(x_296, 1, x_294); return x_296; } } +else +{ +lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +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_297 = lean_ctor_get(x_221, 0); +lean_inc(x_297); +x_298 = lean_ctor_get(x_221, 1); +lean_inc(x_298); +if (lean_is_exclusive(x_221)) { + lean_ctor_release(x_221, 0); + lean_ctor_release(x_221, 1); + x_299 = x_221; +} else { + lean_dec_ref(x_221); + x_299 = lean_box(0); +} +if (lean_is_scalar(x_299)) { + x_300 = lean_alloc_ctor(1, 2, 0); +} else { + x_300 = x_299; +} +lean_ctor_set(x_300, 0, x_297); +lean_ctor_set(x_300, 1, x_298); +return x_300; +} +} } } } @@ -28263,68 +28267,69 @@ return x_78; LEAN_EXPORT lean_object* l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___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) { _start: { -uint8_t x_14; lean_object* x_15; -x_14 = 0; +uint8_t x_14; uint8_t x_15; lean_object* x_16; +x_14 = 1; +x_15 = 0; 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_15 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_14, x_14, x_7, x_8, x_9, x_10, x_11, x_12, x_13); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_14, x_15, x_7, x_8, x_9, x_10, x_11, x_12, x_13); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_7); lean_inc(x_1); -x_17 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getResultUniverse(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_16); -if (lean_obj_tag(x_17) == 0) +x_18 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_getResultUniverse(x_1, x_7, x_8, x_9, x_10, x_11, x_12, x_17); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_18 = lean_ctor_get(x_17, 0); -lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 1); +lean_object* x_19; lean_object* x_20; lean_object* x_21; +x_19 = lean_ctor_get(x_18, 0); lean_inc(x_19); -lean_dec(x_17); +x_20 = lean_ctor_get(x_18, 1); +lean_inc(x_20); +lean_dec(x_18); lean_inc(x_7); -x_20 = l_Lean_Elab_Command_shouldInferResultUniverse(x_18, x_7, x_8, x_9, x_10, x_11, x_12, x_19); -if (lean_obj_tag(x_20) == 0) +x_21 = l_Lean_Elab_Command_shouldInferResultUniverse(x_19, x_7, x_8, x_9, x_10, x_11, x_12, x_20); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_22 = lean_ctor_get(x_21, 0); lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_ctor_get(x_2, 5); +x_23 = lean_ctor_get(x_21, 1); lean_inc(x_23); -x_24 = lean_ctor_get(x_2, 6); +lean_dec(x_21); +x_24 = lean_ctor_get(x_2, 5); lean_inc(x_24); +x_25 = lean_ctor_get(x_2, 6); +lean_inc(x_25); lean_inc(x_6); -lean_inc(x_24); -x_25 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4), 16, 8); -lean_closure_set(x_25, 0, x_24); -lean_closure_set(x_25, 1, x_6); -lean_closure_set(x_25, 2, x_21); -lean_closure_set(x_25, 3, x_2); -lean_closure_set(x_25, 4, x_3); -lean_closure_set(x_25, 5, x_4); -lean_closure_set(x_25, 6, x_5); -lean_closure_set(x_25, 7, x_1); -x_26 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withUsed___rarg(x_23, x_24, x_6, x_25, x_7, x_8, x_9, x_10, x_11, x_12, x_22); -lean_dec(x_23); -return x_26; +lean_inc(x_25); +x_26 = lean_alloc_closure((void*)(l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView___lambda__4), 16, 8); +lean_closure_set(x_26, 0, x_25); +lean_closure_set(x_26, 1, x_6); +lean_closure_set(x_26, 2, x_22); +lean_closure_set(x_26, 3, x_2); +lean_closure_set(x_26, 4, x_3); +lean_closure_set(x_26, 5, x_4); +lean_closure_set(x_26, 6, x_5); +lean_closure_set(x_26, 7, x_1); +x_27 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_withUsed___rarg(x_24, x_25, x_6, x_26, x_7, x_8, x_9, x_10, x_11, x_12, x_23); +lean_dec(x_24); +return x_27; } else { -uint8_t x_27; +uint8_t x_28; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -28337,29 +28342,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_20); -if (x_27 == 0) +x_28 = !lean_is_exclusive(x_21); +if (x_28 == 0) { -return x_20; +return x_21; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_20, 0); -x_29 = lean_ctor_get(x_20, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_21, 0); +x_30 = lean_ctor_get(x_21, 1); +lean_inc(x_30); lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_20); -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_dec(x_21); +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_31; +uint8_t x_32; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -28372,29 +28377,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_17); -if (x_31 == 0) +x_32 = !lean_is_exclusive(x_18); +if (x_32 == 0) { -return x_17; +return x_18; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_17, 0); -x_33 = lean_ctor_get(x_17, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_18, 0); +x_34 = lean_ctor_get(x_18, 1); +lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); -lean_dec(x_17); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +lean_dec(x_18); +x_35 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +return x_35; } } } else { -uint8_t x_35; +uint8_t x_36; lean_dec(x_12); lean_dec(x_11); lean_dec(x_10); @@ -28407,23 +28412,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_35 = !lean_is_exclusive(x_15); -if (x_35 == 0) +x_36 = !lean_is_exclusive(x_16); +if (x_36 == 0) { -return x_15; +return x_16; } else { -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_15, 0); -x_37 = lean_ctor_get(x_15, 1); +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_inc(x_38); lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_15); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_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; } } } @@ -29130,218 +29135,219 @@ return x_10; LEAN_EXPORT lean_object* l_Lean_Elab_Command_elabStructure___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, 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) { _start: { -uint8_t x_20; lean_object* x_21; -x_20 = 0; +uint8_t x_20; uint8_t x_21; lean_object* x_22; +x_20 = 1; +x_21 = 0; 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_21 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_13, x_14, x_15, x_16, x_17, x_18, x_19); -if (lean_obj_tag(x_21) == 0) +x_22 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_21, x_13, x_14, x_15, x_16, x_17, x_18, x_19); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_13); -x_23 = l_Lean_Elab_Term_addAutoBoundImplicits(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_22); -if (lean_obj_tag(x_23) == 0) +x_24 = l_Lean_Elab_Term_addAutoBoundImplicits(x_12, x_13, x_14, x_15, x_16, x_17, x_18, x_23); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -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; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_Elab_Term_getLevelNames___rarg(x_14, x_15, x_16, x_17, x_18, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_Elab_Term_getLevelNames___rarg(x_14, x_15, x_16, x_17, x_18, x_26); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); lean_inc(x_5); -x_29 = lean_alloc_ctor(0, 11, 1); -lean_ctor_set(x_29, 0, x_1); -lean_ctor_set(x_29, 1, x_2); -lean_ctor_set(x_29, 2, x_3); -lean_ctor_set(x_29, 3, x_27); -lean_ctor_set(x_29, 4, x_5); -lean_ctor_set(x_29, 5, x_6); -lean_ctor_set(x_29, 6, x_24); -lean_ctor_set(x_29, 7, x_7); -lean_ctor_set(x_29, 8, x_8); -lean_ctor_set(x_29, 9, x_9); -lean_ctor_set(x_29, 10, x_10); -lean_ctor_set_uint8(x_29, sizeof(void*)*11, x_4); +x_30 = lean_alloc_ctor(0, 11, 1); +lean_ctor_set(x_30, 0, x_1); +lean_ctor_set(x_30, 1, x_2); +lean_ctor_set(x_30, 2, x_3); +lean_ctor_set(x_30, 3, x_28); +lean_ctor_set(x_30, 4, x_5); +lean_ctor_set(x_30, 5, x_6); +lean_ctor_set(x_30, 6, x_25); +lean_ctor_set(x_30, 7, x_7); +lean_ctor_set(x_30, 8, x_8); +lean_ctor_set(x_30, 9, x_9); +lean_ctor_set(x_30, 10, x_10); +lean_ctor_set_uint8(x_30, sizeof(void*)*11, x_4); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); -x_30 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView(x_29, x_13, x_14, x_15, x_16, x_17, x_18, x_28); -if (lean_obj_tag(x_30) == 0) +x_31 = l___private_Lean_Elab_Structure_0__Lean_Elab_Command_elabStructureView(x_30, x_13, x_14, x_15, x_16, x_17, x_18, x_29); +if (lean_obj_tag(x_31) == 0) { if (x_11 == 0) { -lean_object* x_31; lean_object* x_32; -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -lean_dec(x_30); +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_31, 1); +lean_inc(x_32); +lean_dec(x_31); lean_inc(x_18); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_5); -x_32 = l_Lean_Meta_mkSizeOfInstances(x_5, x_15, x_16, x_17, x_18, x_31); -if (lean_obj_tag(x_32) == 0) +x_33 = l_Lean_Meta_mkSizeOfInstances(x_5, x_15, x_16, x_17, x_18, x_32); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_32, 1); -lean_inc(x_33); -lean_dec(x_32); +lean_object* x_34; lean_object* x_35; +x_34 = lean_ctor_get(x_33, 1); +lean_inc(x_34); +lean_dec(x_33); lean_inc(x_5); -x_34 = l_Lean_Meta_mkInjectiveTheorems(x_5, x_15, x_16, x_17, x_18, x_33); -if (lean_obj_tag(x_34) == 0) +x_35 = l_Lean_Meta_mkInjectiveTheorems(x_5, x_15, x_16, x_17, x_18, x_34); +if (lean_obj_tag(x_35) == 0) { -uint8_t x_35; -x_35 = !lean_is_exclusive(x_34); -if (x_35 == 0) +uint8_t x_36; +x_36 = !lean_is_exclusive(x_35); +if (x_36 == 0) { -lean_object* x_36; -x_36 = lean_ctor_get(x_34, 0); -lean_dec(x_36); -lean_ctor_set(x_34, 0, x_5); -return x_34; +lean_object* x_37; +x_37 = lean_ctor_get(x_35, 0); +lean_dec(x_37); +lean_ctor_set(x_35, 0, x_5); +return x_35; } else { -lean_object* x_37; lean_object* x_38; -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); -lean_dec(x_34); -x_38 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_38, 0, x_5); -lean_ctor_set(x_38, 1, x_37); -return x_38; +lean_object* x_38; lean_object* x_39; +x_38 = lean_ctor_get(x_35, 1); +lean_inc(x_38); +lean_dec(x_35); +x_39 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_39, 0, x_5); +lean_ctor_set(x_39, 1, x_38); +return x_39; } } else { -uint8_t x_39; +uint8_t x_40; lean_dec(x_5); -x_39 = !lean_is_exclusive(x_34); -if (x_39 == 0) +x_40 = !lean_is_exclusive(x_35); +if (x_40 == 0) { -return x_34; +return x_35; } else { -lean_object* x_40; lean_object* x_41; lean_object* x_42; -x_40 = lean_ctor_get(x_34, 0); -x_41 = lean_ctor_get(x_34, 1); +lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_41 = lean_ctor_get(x_35, 0); +x_42 = lean_ctor_get(x_35, 1); +lean_inc(x_42); lean_inc(x_41); -lean_inc(x_40); -lean_dec(x_34); -x_42 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_42, 0, x_40); -lean_ctor_set(x_42, 1, x_41); -return x_42; +lean_dec(x_35); +x_43 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_43, 0, x_41); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } } else { -uint8_t x_43; +uint8_t x_44; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_5); -x_43 = !lean_is_exclusive(x_32); -if (x_43 == 0) +x_44 = !lean_is_exclusive(x_33); +if (x_44 == 0) { -return x_32; +return x_33; } else { -lean_object* x_44; lean_object* x_45; lean_object* x_46; -x_44 = lean_ctor_get(x_32, 0); -x_45 = lean_ctor_get(x_32, 1); +lean_object* x_45; lean_object* x_46; lean_object* x_47; +x_45 = lean_ctor_get(x_33, 0); +x_46 = lean_ctor_get(x_33, 1); +lean_inc(x_46); lean_inc(x_45); -lean_inc(x_44); -lean_dec(x_32); -x_46 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_46, 0, x_44); -lean_ctor_set(x_46, 1, x_45); -return x_46; +lean_dec(x_33); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_45); +lean_ctor_set(x_47, 1, x_46); +return x_47; } } } else { -uint8_t x_47; +uint8_t x_48; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -x_47 = !lean_is_exclusive(x_30); -if (x_47 == 0) +x_48 = !lean_is_exclusive(x_31); +if (x_48 == 0) { -lean_object* x_48; -x_48 = lean_ctor_get(x_30, 0); -lean_dec(x_48); -lean_ctor_set(x_30, 0, x_5); -return x_30; +lean_object* x_49; +x_49 = lean_ctor_get(x_31, 0); +lean_dec(x_49); +lean_ctor_set(x_31, 0, x_5); +return x_31; } else { -lean_object* x_49; lean_object* x_50; -x_49 = lean_ctor_get(x_30, 1); -lean_inc(x_49); -lean_dec(x_30); -x_50 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_50, 0, x_5); -lean_ctor_set(x_50, 1, x_49); -return x_50; +lean_object* x_50; lean_object* x_51; +x_50 = lean_ctor_get(x_31, 1); +lean_inc(x_50); +lean_dec(x_31); +x_51 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_51, 0, x_5); +lean_ctor_set(x_51, 1, x_50); +return x_51; } } } else { -uint8_t x_51; +uint8_t x_52; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); lean_dec(x_5); -x_51 = !lean_is_exclusive(x_30); -if (x_51 == 0) +x_52 = !lean_is_exclusive(x_31); +if (x_52 == 0) { -return x_30; +return x_31; } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_30, 0); -x_53 = lean_ctor_get(x_30, 1); +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_inc(x_52); -lean_dec(x_30); -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_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; } } } else { -uint8_t x_55; +uint8_t x_56; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -29357,29 +29363,29 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_55 = !lean_is_exclusive(x_23); -if (x_55 == 0) +x_56 = !lean_is_exclusive(x_24); +if (x_56 == 0) { -return x_23; +return x_24; } else { -lean_object* x_56; lean_object* x_57; lean_object* x_58; -x_56 = lean_ctor_get(x_23, 0); -x_57 = lean_ctor_get(x_23, 1); +lean_object* x_57; lean_object* x_58; lean_object* x_59; +x_57 = lean_ctor_get(x_24, 0); +x_58 = lean_ctor_get(x_24, 1); +lean_inc(x_58); lean_inc(x_57); -lean_inc(x_56); -lean_dec(x_23); -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_dec(x_24); +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_59; +uint8_t x_60; lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); @@ -29396,23 +29402,23 @@ lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_59 = !lean_is_exclusive(x_21); -if (x_59 == 0) +x_60 = !lean_is_exclusive(x_22); +if (x_60 == 0) { -return x_21; +return x_22; } else { -lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_60 = lean_ctor_get(x_21, 0); -x_61 = lean_ctor_get(x_21, 1); +lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_61 = lean_ctor_get(x_22, 0); +x_62 = lean_ctor_get(x_22, 1); +lean_inc(x_62); lean_inc(x_61); -lean_inc(x_60); -lean_dec(x_21); -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_dec(x_22); +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; } } } diff --git a/stage0/stdlib/Lean/Elab/SyntheticMVars.c b/stage0/stdlib/Lean/Elab/SyntheticMVars.c index ff49f5b38da9..de543f9cd038 100644 --- a/stage0/stdlib/Lean/Elab/SyntheticMVars.c +++ b/stage0/stdlib/Lean/Elab/SyntheticMVars.c @@ -14,20 +14,23 @@ extern "C" { #endif LEAN_EXPORT lean_object* l_Lean_getExprMVarAssignment_x3f___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashSetImp_contains___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___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_PostponeBehavior_noConfusion___rarg(uint8_t, uint8_t, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___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_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVar___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__7; LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___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_mkHashSet___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__20; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__2; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__6; LEAN_EXPORT lean_object* l_Lean_RBNode_del___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_markAsResolved___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -38,6 +41,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic___lambda__3(lean_object*, le LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesizeLight___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_instantiateMVarDeclMVars___spec__1(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Elab_Term_instInhabitedPostponeBehavior; lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___closed__2; extern lean_object* l_Lean_maxRecDepthErrorMessage; @@ -55,10 +59,13 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Elab_Term_runTa lean_object* l_Lean_Elab_Term_elabTerm(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16; size_t lean_hashset_mk_idx(lean_object*, uint64_t); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_toCtorIdx___boxed(lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___spec__1___closed__1; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__2___closed__2; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___spec__2___closed__3; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_instReprPostponeBehavior; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_runTactic___spec__8(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___closed__2; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___closed__3; @@ -72,14 +79,11 @@ lean_object* l_Lean_mkHashSetImp___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio___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_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__5; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__17; lean_object* lean_mk_array(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12; uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesizeLight(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__11; -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3; lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__11; lean_object* l_Lean_replaceRef(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__3; lean_object* l_Lean_Elab_Term_getMVarDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -88,16 +92,20 @@ LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_S 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_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19; LEAN_EXPORT lean_object* l_Lean_mkHashSet___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstancesPriorities___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___spec__1___rarg___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__22; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___closed__2; LEAN_EXPORT lean_object* l_List_elem___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__3___boxed(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__1; LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__16___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__13(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* l_Lean_Elab_logException___at_Lean_Elab_Term_exceptionToSorry___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -113,16 +121,18 @@ static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resum static lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_Term_runTactic___spec__2___closed__5; lean_object* l_Lean_stringToMessageData(lean_object*); lean_object* l_Lean_Expr_mvar___override(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg(lean_object*, lean_object*, lean_object*, uint8_t); lean_object* l_Lean_RBNode_balRight___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_occursCheck___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_Term_runTactic___spec__2___closed__2; lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20; LEAN_EXPORT lean_object* l_Lean_HashSetImp_expand___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__5(lean_object*, lean_object*); lean_object* l_instBEqProd___rarg(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*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__9; +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars_loop___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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*, lean_object*, lean_object*, lean_object*); @@ -131,11 +141,13 @@ static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synth static lean_object* l_Lean_Elab_Term_runTactic___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___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_Term_synthesizeSyntheticMVarsNoPostponing___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_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_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*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4; LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_processPostponedUniverseContraints(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_List_elem___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101____boxed(lean_object*, lean_object*); lean_object* l_List_appendTR___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_instInhabitedInfoTree; @@ -144,23 +156,25 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesizeLight___rarg___boxed(lea lean_object* l_Lean_Elab_Term_withSavedContext___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_Term_runTactic___spec__2___closed__7; 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_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13; +static lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_appendTrees___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_getSomeSyntheticMVarsRef(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_mkLeveErrorMessageCore(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__14; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6; LEAN_EXPORT lean_object* l_Lean_occursCheck_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultLoop(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Term_runTactic___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__13; LEAN_EXPORT lean_object* l_Lean_Elab_Term_runPendingTacticsAt(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); +lean_object* lean_nat_to_int(lean_object*); static lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___closed__1; lean_object* l_Lean_MVarId_getType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_forallMetaTelescopeReducingAux(lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -172,11 +186,13 @@ lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingInstancesStep___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_runTactic___spec__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_Term_runTactic___spec__2___closed__4; +static lean_object* l_Lean_Elab_Term_instReprPostponeBehavior___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_elabTerm___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986_(uint8_t, lean_object*); lean_object* l_Lean_Elab_goalsToMessageData(lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___boxed(lean_object*, lean_object*, lean_object*); @@ -186,15 +202,17 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic(lean_object*, lean_object*, LEAN_EXPORT lean_object* l_Lean_instantiateLCtxMVars___at_Lean_Elab_Term_runTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__7(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_markAsResolved___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___closed__2; extern uint8_t l_Lean_instInhabitedBinderInfo; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__16; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(uint8_t, uint8_t); uint8_t l_List_isEmpty___rarg(lean_object*); lean_object* l_Lean_Elab_Term_traceAtCmdPos(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_LocalContext_mkLetDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t); @@ -212,19 +230,25 @@ static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synth static lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstances___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingInstances(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_runTactic___spec__7(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5; extern lean_object* l_Lean_Meta_instInhabitedDefaultInstances; LEAN_EXPORT lean_object* l_Lean_PersistentArray_forInAux___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__11___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_instInhabitedExpr; extern lean_object* l_Lean_Meta_defaultInstanceExtension; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___lambda__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizePending___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_runTactic___lambda__3___closed__2; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_ofBool___boxed(lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__18; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___lambda__1___closed__4; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14; lean_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSomeUsingDefaultPrio_visit(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_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__3; uint8_t lean_name_eq(lean_object*, lean_object*); @@ -239,7 +263,9 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(l static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVar___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__8; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7; static lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Term_runTactic___spec__5___closed__1; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17; LEAN_EXPORT lean_object* l_Lean_HashSetImp_insert___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__4(lean_object*, lean_object*); lean_object* l_Lean_MetavarContext_getExprAssignmentCore_x3f(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); @@ -250,19 +276,22 @@ lean_object* l_Lean_Elab_Term_throwTypeMismatchError___rarg(lean_object*, lean_o LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_occursCheck_visitMVar___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_getSomeSyntheticMVarsRef___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizePendingInstMVar_x27(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_reportUnsolvedGoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_LocalContext_foldlM___at_Lean_Elab_Term_runTactic___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize(lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___closed__3; -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___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_List_foldl___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__7___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__8(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__13___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15; static lean_object* l_Std_Range_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__15___closed__1; lean_object* lean_usize_to_nat(size_t); lean_object* l_Lean_MVarId_assign___at_Lean_Elab_Term_exprToSyntax___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -291,7 +320,6 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_ LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1(lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic___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_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVar___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -304,14 +332,17 @@ LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVa LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_Term_runTactic___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2; 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*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVar___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___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_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l___private_Lean_Expr_0__Lean_beqBinderInfo____x40_Lean_Expr___hyg_399_(uint8_t, uint8_t); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); lean_object* l_Lean_Meta_coerce_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVar___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -320,26 +351,30 @@ lean_object* l_Lean_indentExpr(lean_object*); lean_object* l_Lean_Meta_isExprDefEqGuarded(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__21; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_runTactic___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*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18; LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__13; lean_object* l_Lean_getDelayedMVarRoot___at_Lean_Elab_Term_isLetRecAuxMVar___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Repr_addAppParen(lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3; LEAN_EXPORT lean_object* l_Lean_commitWhen___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizePendingInstMVar_x27___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTermAndSynthesize(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__10; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1___boxed(lean_object*); lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumeElabTerm___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_PersistentArray_forInAux___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__11___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_withoutAutoBoundImplicit___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSomeUsingDefaultPrio(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_SyntheticMVars___hyg_6641____closed__8; LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__5; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumeElabTerm(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Elab_postponeExceptionId; @@ -350,6 +385,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic___lambda__3___boxed(lean_obj LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint64_t l_Lean_Level_hash(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__6; lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_getDelayedMVarAssignment_x3f___spec__1(lean_object*, lean_object*); @@ -363,8 +399,10 @@ LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at___private_Lean LEAN_EXPORT lean_object* l_Lean_Elab_Term_runTactic___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstancesPriorities___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__9; uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__5; size_t lean_usize_add(size_t, size_t); lean_object* l_Lean_Meta_getMVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_replace___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__9(lean_object*, lean_object*, lean_object*); @@ -373,6 +411,7 @@ static lean_object* l_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0_ LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runTactic___spec__10(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__2___closed__3; lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio(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_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___spec__1___closed__1; lean_object* lean_array_uget(lean_object*, size_t); @@ -381,20 +420,23 @@ lean_object* l_Lean_Elab_Tactic_mkInitialTacticInfo(lean_object*, lean_object*, lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Elab_assignInfoHoleId___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_markAsResolved___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars_loop(uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint8_t l_Lean_Elab_Term_PostponeBehavior_ofBool(uint8_t); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_getSyntheticMVarDecl_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_shift_left(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_getDefaultInstancesPriorities___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_occursCheck_visitMVar___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_resumePostponed___spec__3___closed__1; +LEAN_EXPORT lean_object* l_Lean_Elab_Term_instBEqPostponeBehavior; 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*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_markAsResolved___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__2___closed__5; +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__2; +static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio___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 lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___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_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__7; static lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___spec__2___closed__1; lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -415,7 +457,6 @@ lean_object* lean_expr_dbg_to_string(lean_object*); uint8_t l_Lean_HashSetImp_contains___at_Lean_CollectMVars_visit___spec__1(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); uint8_t l_Lean_Exception_isRuntime(lean_object*); -static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__8; uint8_t l_Lean_Level_normLtAux(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_RBNode_find___at_Lean_Meta_addDefaultInstanceEntry___spec__3(lean_object*, lean_object*); lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -430,6 +471,7 @@ lean_object* l_Lean_MVarId_withContext___at_Lean_Elab_Tactic_run___spec__1___rar lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_getPostponed___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingInstancesStep(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_toCtorIdx(uint8_t); LEAN_EXPORT lean_object* l_Lean_PersistentArray_forIn___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefault___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_List_filterAuxM___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___spec__1___lambda__2___closed__6; @@ -437,6 +479,7 @@ LEAN_EXPORT lean_object* l_Lean_instantiateMVarDeclMVars___at_Lean_Elab_Term_run static lean_object* l_Lean_Elab_Term_reportStuckSyntheticMVar___lambda__1___closed__1; lean_object* l_Lean_HashSetImp_insert___at_Lean_CollectMVars_visit___spec__3(lean_object*, lean_object*); size_t lean_usize_land(size_t, size_t); +static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__15; static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__2___closed__4; LEAN_EXPORT lean_object* l_panic___at___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizePendingInstMVar___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___lambda__1___closed__11; @@ -12057,347 +12100,1207 @@ lean_dec(x_1); return x_9; } } +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_toCtorIdx(uint8_t x_1) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_2; +x_2 = lean_unsigned_to_nat(0u); +return x_2; +} +case 1: +{ +lean_object* x_3; +x_3 = lean_unsigned_to_nat(1u); +return x_3; +} +default: +{ +lean_object* x_4; +x_4 = lean_unsigned_to_nat(2u); +return x_4; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_toCtorIdx___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; lean_object* x_3; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Lean_Elab_Term_PostponeBehavior_toCtorIdx(x_2); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1(lean_object* x_1) { +_start: +{ +lean_inc(x_1); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1___boxed), 1, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg(uint8_t x_1, uint8_t x_2, lean_object* x_3) { +_start: +{ +lean_object* x_4; +x_4 = l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1; +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___boxed), 3, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1___boxed(lean_object* x_1) { +_start: +{ +lean_object* x_2; +x_2 = l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___lambda__1(x_1); +lean_dec(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +_start: +{ +uint8_t x_4; uint8_t x_5; lean_object* x_6; +x_4 = lean_unbox(x_1); +lean_dec(x_1); +x_5 = lean_unbox(x_2); +lean_dec(x_2); +x_6 = l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg(x_4, x_5, x_3); +return x_6; +} +} +static uint8_t _init_l_Lean_Elab_Term_instInhabitedPostponeBehavior() { +_start: +{ +uint8_t x_1; +x_1 = 0; +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Elab.Term.PostponeBehavior.yes", 35); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____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_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(2u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = lean_unsigned_to_nat(1u); +x_2 = lean_nat_to_int(x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Elab.Term.PostponeBehavior.no", 34); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10; +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_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10; +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_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("Lean.Elab.Term.PostponeBehavior.partial", 39); +return x_1; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15; +x_2 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17; +x_2 = 0; +x_3 = lean_alloc_ctor(6, 1, 1); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set_uint8(x_3, sizeof(void*)*1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19() { +_start: +{ +lean_object* x_1; lean_object* x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6; +x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16; +x_3 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_3, 0, x_1); +lean_ctor_set(x_3, 1, x_2); +return x_3; +} +} +static lean_object* _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20() { +_start: +{ +lean_object* x_1; uint8_t x_2; lean_object* x_3; +x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19; +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_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986_(uint8_t x_1, lean_object* x_2) { +_start: +{ +switch (x_1) { +case 0: +{ +lean_object* x_3; uint8_t x_4; +x_3 = lean_unsigned_to_nat(1024u); +x_4 = lean_nat_dec_le(x_3, x_2); +if (x_4 == 0) +{ +lean_object* x_5; lean_object* x_6; +x_5 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5; +x_6 = l_Repr_addAppParen(x_5, x_2); +return x_6; +} +else +{ +lean_object* x_7; lean_object* x_8; +x_7 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8; +x_8 = l_Repr_addAppParen(x_7, x_2); +return x_8; +} +} +case 1: +{ +lean_object* x_9; uint8_t x_10; +x_9 = lean_unsigned_to_nat(1024u); +x_10 = lean_nat_dec_le(x_9, x_2); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; +x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12; +x_12 = l_Repr_addAppParen(x_11, x_2); +return x_12; +} +else +{ +lean_object* x_13; lean_object* x_14; +x_13 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14; +x_14 = l_Repr_addAppParen(x_13, x_2); +return x_14; +} +} +default: +{ +lean_object* x_15; uint8_t x_16; +x_15 = lean_unsigned_to_nat(1024u); +x_16 = lean_nat_dec_le(x_15, x_2); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18; +x_18 = l_Repr_addAppParen(x_17, x_2); +return x_18; +} +else +{ +lean_object* x_19; lean_object* x_20; +x_19 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20; +x_20 = l_Repr_addAppParen(x_19, x_2); +return x_20; +} +} +} +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = lean_unbox(x_1); +lean_dec(x_1); +x_4 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986_(x_3, x_2); +lean_dec(x_2); +return x_4; +} +} +static lean_object* _init_l_Lean_Elab_Term_instReprPostponeBehavior___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_instReprPostponeBehavior() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Term_instReprPostponeBehavior___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(uint8_t x_1, uint8_t x_2) { +_start: +{ +lean_object* x_3; lean_object* x_4; uint8_t x_5; +x_3 = l_Lean_Elab_Term_PostponeBehavior_toCtorIdx(x_1); +x_4 = l_Lean_Elab_Term_PostponeBehavior_toCtorIdx(x_2); +x_5 = lean_nat_dec_eq(x_3, x_4); +lean_dec(x_4); +lean_dec(x_3); +return x_5; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101____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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_3, x_4); +x_6 = lean_box(x_5); +return x_6; +} +} +static lean_object* _init_l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101____boxed), 2, 0); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_instBEqPostponeBehavior() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1; +return x_1; +} +} +LEAN_EXPORT uint8_t l_Lean_Elab_Term_PostponeBehavior_ofBool(uint8_t x_1) { +_start: +{ +if (x_1 == 0) +{ +uint8_t x_2; +x_2 = 1; +return x_2; +} +else +{ +uint8_t x_3; +x_3 = 0; +return x_3; +} +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_PostponeBehavior_ofBool___boxed(lean_object* x_1) { +_start: +{ +uint8_t x_2; uint8_t x_3; lean_object* x_4; +x_2 = lean_unbox(x_1); +lean_dec(x_1); +x_3 = l_Lean_Elab_Term_PostponeBehavior_ofBool(x_2); +x_4 = lean_box(x_3); +return x_4; +} +} static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1() { _start: { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_maxRecDepthErrorMessage; -x_2 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_maxRecDepthErrorMessage; +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_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1; +x_2 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_2, 0, x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_9 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2; +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_1); +lean_ctor_set(x_10, 1, x_9); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_8); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars_loop(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { +_start: +{ +lean_object* x_11; +lean_dec(x_3); +lean_inc(x_9); +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_getSomeSyntheticMVarsRef(x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; lean_object* x_13; uint8_t 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_is_exclusive(x_8); +if (x_14 == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; uint8_t x_27; lean_object* x_28; uint8_t x_29; +x_15 = lean_ctor_get(x_8, 0); +x_16 = lean_ctor_get(x_8, 1); +x_17 = lean_ctor_get(x_8, 2); +x_18 = lean_ctor_get(x_8, 3); +x_19 = lean_ctor_get(x_8, 4); +x_20 = lean_ctor_get(x_8, 5); +x_21 = lean_ctor_get(x_8, 6); +x_22 = lean_ctor_get(x_8, 7); +x_23 = lean_ctor_get(x_8, 8); +x_24 = lean_ctor_get(x_8, 9); +x_25 = lean_ctor_get(x_8, 10); +x_26 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +x_27 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); +x_28 = l_Lean_replaceRef(x_12, x_20); +lean_dec(x_20); +lean_dec(x_12); +lean_inc(x_25); +lean_inc(x_24); +lean_inc(x_23); +lean_inc(x_22); +lean_inc(x_21); +lean_inc(x_28); +lean_inc(x_19); +lean_inc(x_18); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_ctor_set(x_8, 5, x_28); +x_29 = lean_nat_dec_eq(x_18, x_19); +if (x_29 == 0) +{ +lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; +lean_dec(x_8); +x_30 = lean_unsigned_to_nat(1u); +x_31 = lean_nat_add(x_18, x_30); +lean_dec(x_18); +x_32 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_32, 0, x_15); +lean_ctor_set(x_32, 1, x_16); +lean_ctor_set(x_32, 2, x_17); +lean_ctor_set(x_32, 3, x_31); +lean_ctor_set(x_32, 4, x_19); +lean_ctor_set(x_32, 5, x_28); +lean_ctor_set(x_32, 6, x_21); +lean_ctor_set(x_32, 7, x_22); +lean_ctor_set(x_32, 8, x_23); +lean_ctor_set(x_32, 9, x_24); +lean_ctor_set(x_32, 10, x_25); +lean_ctor_set_uint8(x_32, sizeof(void*)*11, x_26); +lean_ctor_set_uint8(x_32, sizeof(void*)*11 + 1, x_27); +x_33 = lean_st_ref_get(x_5, x_13); +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; +x_35 = lean_ctor_get(x_33, 0); +x_36 = lean_ctor_get(x_33, 1); +x_37 = lean_ctor_get(x_35, 2); +lean_inc(x_37); +lean_dec(x_35); +x_38 = l_List_isEmpty___rarg(x_37); +lean_dec(x_37); +if (x_38 == 0) +{ +uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +lean_free_object(x_33); +x_39 = 0; +x_40 = lean_box(x_39); +x_41 = lean_box(x_39); +x_42 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_42, 0, x_40); +lean_closure_set(x_42, 1, x_41); +lean_inc(x_9); +lean_inc(x_32); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_43 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_39, x_39, x_4, x_5, x_6, x_7, x_32, x_9, x_36); +if (lean_obj_tag(x_43) == 0) +{ +lean_object* x_44; uint8_t x_45; +x_44 = lean_ctor_get(x_43, 0); +lean_inc(x_44); +x_45 = lean_unbox(x_44); +lean_dec(x_44); +if (x_45 == 0) +{ +uint8_t x_46; +x_46 = !lean_is_exclusive(x_43); +if (x_46 == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; uint8_t x_50; +x_47 = lean_ctor_get(x_43, 1); +x_48 = lean_ctor_get(x_43, 0); +lean_dec(x_48); +x_49 = 0; +x_50 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_49); +if (x_50 == 0) +{ +uint8_t x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; +lean_free_object(x_43); +x_51 = 1; +x_52 = lean_box(x_51); +x_53 = lean_box(x_39); +x_54 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_54, 0, x_52); +lean_closure_set(x_54, 1, x_53); +lean_inc(x_9); +lean_inc(x_32); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_55 = l_Lean_Elab_Term_withoutPostponing___rarg(x_54, x_4, x_5, x_6, x_7, x_32, x_9, x_47); +if (lean_obj_tag(x_55) == 0) +{ +lean_object* x_56; uint8_t x_57; +x_56 = lean_ctor_get(x_55, 0); +lean_inc(x_56); +x_57 = lean_unbox(x_56); +lean_dec(x_56); +if (x_57 == 0) +{ +lean_object* x_58; lean_object* x_59; +x_58 = lean_ctor_get(x_55, 1); +lean_inc(x_58); +lean_dec(x_55); +lean_inc(x_9); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_59 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_32, x_9, x_58); +if (lean_obj_tag(x_59) == 0) +{ +lean_object* x_60; uint8_t x_61; +x_60 = lean_ctor_get(x_59, 0); +lean_inc(x_60); +x_61 = lean_unbox(x_60); +lean_dec(x_60); +if (x_61 == 0) +{ +lean_object* x_62; lean_object* x_63; +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +lean_inc(x_9); +lean_inc(x_32); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_63 = l_Lean_Elab_Term_withoutPostponing___rarg(x_42, x_4, x_5, x_6, x_7, x_32, x_9, x_62); +if (lean_obj_tag(x_63) == 0) +{ +lean_object* x_64; uint8_t x_65; +x_64 = lean_ctor_get(x_63, 0); +lean_inc(x_64); +x_65 = lean_unbox(x_64); +lean_dec(x_64); +if (x_65 == 0) +{ +lean_object* x_66; lean_object* x_67; +x_66 = lean_ctor_get(x_63, 1); +lean_inc(x_66); +lean_dec(x_63); +lean_inc(x_9); +lean_inc(x_32); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +x_67 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_39, x_51, x_4, x_5, x_6, x_7, x_32, x_9, x_66); +if (lean_obj_tag(x_67) == 0) +{ +lean_object* x_68; uint8_t x_69; +x_68 = lean_ctor_get(x_67, 0); +lean_inc(x_68); +x_69 = lean_unbox(x_68); +lean_dec(x_68); +if (x_69 == 0) +{ +uint8_t x_70; +x_70 = !lean_is_exclusive(x_67); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; uint8_t x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_67, 1); +x_72 = lean_ctor_get(x_67, 0); +lean_dec(x_72); +x_73 = 1; +x_74 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_73); +if (x_74 == 0) +{ +lean_object* x_75; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_75 = lean_box(0); +lean_ctor_set(x_67, 0, x_75); +return x_67; +} +else +{ +lean_object* x_76; +lean_free_object(x_67); +x_76 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_71); +return x_76; +} +} +else +{ +lean_object* x_77; uint8_t x_78; uint8_t x_79; +x_77 = lean_ctor_get(x_67, 1); +lean_inc(x_77); +lean_dec(x_67); +x_78 = 1; +x_79 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_78); +if (x_79 == 0) +{ +lean_object* x_80; lean_object* x_81; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_80 = lean_box(0); +x_81 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_81, 0, x_80); +lean_ctor_set(x_81, 1, x_77); +return x_81; +} +else +{ +lean_object* x_82; +x_82 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_77); +return x_82; +} +} +} +else +{ +lean_object* x_83; lean_object* x_84; +x_83 = lean_ctor_get(x_67, 1); +lean_inc(x_83); +lean_dec(x_67); +x_84 = lean_box(0); +x_3 = x_84; +x_8 = x_32; +x_10 = x_83; +goto _start; +} +} +else +{ +uint8_t x_86; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_86 = !lean_is_exclusive(x_67); +if (x_86 == 0) +{ +return x_67; } +else +{ +lean_object* x_87; lean_object* x_88; lean_object* x_89; +x_87 = lean_ctor_get(x_67, 0); +x_88 = lean_ctor_get(x_67, 1); +lean_inc(x_88); +lean_inc(x_87); +lean_dec(x_67); +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; } -static lean_object* _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2() { -_start: +} +} +else { -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1; -x_2 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_2, 0, x_1); -return x_2; +lean_object* x_90; lean_object* x_91; +x_90 = lean_ctor_get(x_63, 1); +lean_inc(x_90); +lean_dec(x_63); +x_91 = lean_box(0); +x_3 = x_91; +x_8 = x_32; +x_10 = x_90; +goto _start; } } -LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: +else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2; -x_10 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_10, 0, x_1); -lean_ctor_set(x_10, 1, x_9); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_8); -return x_11; +uint8_t x_93; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_93 = !lean_is_exclusive(x_63); +if (x_93 == 0) +{ +return x_63; } +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_63, 0); +x_95 = lean_ctor_get(x_63, 1); +lean_inc(x_95); +lean_inc(x_94); +lean_dec(x_63); +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_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVars_loop(uint8_t x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: +} +} +else { -lean_object* x_11; -lean_dec(x_3); -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_getSomeSyntheticMVarsRef(x_4, x_5, x_6, x_7, x_8, x_9, x_10); -if (lean_obj_tag(x_11) == 0) +lean_object* x_97; lean_object* x_98; +lean_dec(x_42); +x_97 = lean_ctor_get(x_59, 1); +lean_inc(x_97); +lean_dec(x_59); +x_98 = lean_box(0); +x_3 = x_98; +x_8 = x_32; +x_10 = x_97; +goto _start; +} +} +else { -lean_object* x_12; lean_object* x_13; uint8_t 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_is_exclusive(x_8); -if (x_14 == 0) +uint8_t x_100; +lean_dec(x_42); +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_100 = !lean_is_exclusive(x_59); +if (x_100 == 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; uint8_t x_26; uint8_t x_27; lean_object* x_28; uint8_t x_29; -x_15 = lean_ctor_get(x_8, 0); -x_16 = lean_ctor_get(x_8, 1); -x_17 = lean_ctor_get(x_8, 2); -x_18 = lean_ctor_get(x_8, 3); -x_19 = lean_ctor_get(x_8, 4); -x_20 = lean_ctor_get(x_8, 5); -x_21 = lean_ctor_get(x_8, 6); -x_22 = lean_ctor_get(x_8, 7); -x_23 = lean_ctor_get(x_8, 8); -x_24 = lean_ctor_get(x_8, 9); -x_25 = lean_ctor_get(x_8, 10); -x_26 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -x_27 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); -x_28 = l_Lean_replaceRef(x_12, x_20); -lean_dec(x_20); -lean_dec(x_12); -lean_inc(x_25); -lean_inc(x_24); -lean_inc(x_23); -lean_inc(x_22); -lean_inc(x_21); -lean_inc(x_28); -lean_inc(x_19); -lean_inc(x_18); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_ctor_set(x_8, 5, x_28); -x_29 = lean_nat_dec_eq(x_18, x_19); -if (x_29 == 0) +return x_59; +} +else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -lean_dec(x_8); -x_30 = lean_unsigned_to_nat(1u); -x_31 = lean_nat_add(x_18, x_30); -lean_dec(x_18); -x_32 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_32, 0, x_15); -lean_ctor_set(x_32, 1, x_16); -lean_ctor_set(x_32, 2, x_17); -lean_ctor_set(x_32, 3, x_31); -lean_ctor_set(x_32, 4, x_19); -lean_ctor_set(x_32, 5, x_28); -lean_ctor_set(x_32, 6, x_21); -lean_ctor_set(x_32, 7, x_22); -lean_ctor_set(x_32, 8, x_23); -lean_ctor_set(x_32, 9, x_24); -lean_ctor_set(x_32, 10, x_25); -lean_ctor_set_uint8(x_32, sizeof(void*)*11, x_26); -lean_ctor_set_uint8(x_32, sizeof(void*)*11 + 1, x_27); -x_33 = lean_st_ref_get(x_5, x_13); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +lean_object* x_101; lean_object* x_102; lean_object* x_103; +x_101 = lean_ctor_get(x_59, 0); +x_102 = lean_ctor_get(x_59, 1); +lean_inc(x_102); +lean_inc(x_101); +lean_dec(x_59); +x_103 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_103, 0, x_101); +lean_ctor_set(x_103, 1, x_102); +return x_103; +} +} +} +else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; uint8_t x_38; -x_35 = lean_ctor_get(x_33, 0); -x_36 = lean_ctor_get(x_33, 1); -x_37 = lean_ctor_get(x_35, 2); -lean_inc(x_37); -lean_dec(x_35); -x_38 = l_List_isEmpty___rarg(x_37); -lean_dec(x_37); -if (x_38 == 0) +lean_object* x_104; lean_object* x_105; +lean_dec(x_42); +x_104 = lean_ctor_get(x_55, 1); +lean_inc(x_104); +lean_dec(x_55); +x_105 = lean_box(0); +x_3 = x_105; +x_8 = x_32; +x_10 = x_104; +goto _start; +} +} +else { -uint8_t x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; -lean_free_object(x_33); -x_39 = 0; -x_40 = lean_box(x_39); -x_41 = lean_box(x_39); -x_42 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_42, 0, x_40); -lean_closure_set(x_42, 1, x_41); -lean_inc(x_9); -lean_inc(x_32); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_43 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_39, x_39, x_4, x_5, x_6, x_7, x_32, x_9, x_36); -if (lean_obj_tag(x_43) == 0) +uint8_t x_107; +lean_dec(x_42); +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_107 = !lean_is_exclusive(x_55); +if (x_107 == 0) { -lean_object* x_44; uint8_t x_45; -x_44 = lean_ctor_get(x_43, 0); -lean_inc(x_44); -x_45 = lean_unbox(x_44); -lean_dec(x_44); -if (x_45 == 0) +return x_55; +} +else { -if (x_1 == 0) +lean_object* x_108; lean_object* x_109; lean_object* x_110; +x_108 = lean_ctor_get(x_55, 0); +x_109 = lean_ctor_get(x_55, 1); +lean_inc(x_109); +lean_inc(x_108); +lean_dec(x_55); +x_110 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_110, 0, x_108); +lean_ctor_set(x_110, 1, x_109); +return x_110; +} +} +} +else { -lean_object* x_46; uint8_t x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_46 = lean_ctor_get(x_43, 1); -lean_inc(x_46); +lean_object* x_111; +lean_dec(x_42); +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_111 = lean_box(0); +lean_ctor_set(x_43, 0, x_111); +return x_43; +} +} +else +{ +lean_object* x_112; uint8_t x_113; uint8_t x_114; +x_112 = lean_ctor_get(x_43, 1); +lean_inc(x_112); lean_dec(x_43); -x_47 = 1; -x_48 = lean_box(x_47); -x_49 = lean_box(x_39); -x_50 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_50, 0, x_48); -lean_closure_set(x_50, 1, x_49); +x_113 = 0; +x_114 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_113); +if (x_114 == 0) +{ +uint8_t x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_115 = 1; +x_116 = lean_box(x_115); +x_117 = lean_box(x_39); +x_118 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_118, 0, x_116); +lean_closure_set(x_118, 1, x_117); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_51 = l_Lean_Elab_Term_withoutPostponing___rarg(x_50, x_4, x_5, x_6, x_7, x_32, x_9, x_46); -if (lean_obj_tag(x_51) == 0) +x_119 = l_Lean_Elab_Term_withoutPostponing___rarg(x_118, x_4, x_5, x_6, x_7, x_32, x_9, x_112); +if (lean_obj_tag(x_119) == 0) { -lean_object* x_52; uint8_t x_53; -x_52 = lean_ctor_get(x_51, 0); -lean_inc(x_52); -x_53 = lean_unbox(x_52); -lean_dec(x_52); -if (x_53 == 0) +lean_object* x_120; uint8_t x_121; +x_120 = lean_ctor_get(x_119, 0); +lean_inc(x_120); +x_121 = lean_unbox(x_120); +lean_dec(x_120); +if (x_121 == 0) { -lean_object* x_54; lean_object* x_55; -x_54 = lean_ctor_get(x_51, 1); -lean_inc(x_54); -lean_dec(x_51); +lean_object* x_122; lean_object* x_123; +x_122 = lean_ctor_get(x_119, 1); +lean_inc(x_122); +lean_dec(x_119); lean_inc(x_9); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_55 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_32, x_9, x_54); -if (lean_obj_tag(x_55) == 0) -{ -lean_object* x_56; uint8_t x_57; -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -x_57 = lean_unbox(x_56); -lean_dec(x_56); -if (x_57 == 0) +x_123 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_32, x_9, x_122); +if (lean_obj_tag(x_123) == 0) { -lean_object* x_58; lean_object* x_59; -x_58 = lean_ctor_get(x_55, 1); -lean_inc(x_58); -lean_dec(x_55); +lean_object* x_124; uint8_t x_125; +x_124 = lean_ctor_get(x_123, 0); +lean_inc(x_124); +x_125 = lean_unbox(x_124); +lean_dec(x_124); +if (x_125 == 0) +{ +lean_object* x_126; lean_object* x_127; +x_126 = lean_ctor_get(x_123, 1); +lean_inc(x_126); +lean_dec(x_123); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_59 = l_Lean_Elab_Term_withoutPostponing___rarg(x_42, x_4, x_5, x_6, x_7, x_32, x_9, x_58); -if (lean_obj_tag(x_59) == 0) +x_127 = l_Lean_Elab_Term_withoutPostponing___rarg(x_42, x_4, x_5, x_6, x_7, x_32, x_9, x_126); +if (lean_obj_tag(x_127) == 0) { -lean_object* x_60; uint8_t x_61; -x_60 = lean_ctor_get(x_59, 0); -lean_inc(x_60); -x_61 = lean_unbox(x_60); -lean_dec(x_60); -if (x_61 == 0) +lean_object* x_128; uint8_t x_129; +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); +x_129 = lean_unbox(x_128); +lean_dec(x_128); +if (x_129 == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_ctor_get(x_59, 1); -lean_inc(x_62); -lean_dec(x_59); +lean_object* x_130; lean_object* x_131; +x_130 = lean_ctor_get(x_127, 1); +lean_inc(x_130); +lean_dec(x_127); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_63 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_39, x_47, x_4, x_5, x_6, x_7, x_32, x_9, x_62); -if (lean_obj_tag(x_63) == 0) +x_131 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_39, x_115, x_4, x_5, x_6, x_7, x_32, x_9, x_130); +if (lean_obj_tag(x_131) == 0) { -lean_object* x_64; uint8_t x_65; -x_64 = lean_ctor_get(x_63, 0); -lean_inc(x_64); -x_65 = lean_unbox(x_64); -lean_dec(x_64); -if (x_65 == 0) +lean_object* x_132; uint8_t x_133; +x_132 = lean_ctor_get(x_131, 0); +lean_inc(x_132); +x_133 = lean_unbox(x_132); +lean_dec(x_132); +if (x_133 == 0) { -lean_object* x_66; lean_object* x_67; -x_66 = lean_ctor_get(x_63, 1); -lean_inc(x_66); -lean_dec(x_63); -x_67 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_66); -return x_67; +lean_object* x_134; lean_object* x_135; uint8_t x_136; uint8_t x_137; +x_134 = lean_ctor_get(x_131, 1); +lean_inc(x_134); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_135 = x_131; +} else { + lean_dec_ref(x_131); + x_135 = lean_box(0); +} +x_136 = 1; +x_137 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_136); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_138 = lean_box(0); +if (lean_is_scalar(x_135)) { + x_139 = lean_alloc_ctor(0, 2, 0); +} else { + x_139 = x_135; +} +lean_ctor_set(x_139, 0, x_138); +lean_ctor_set(x_139, 1, x_134); +return x_139; } else { -lean_object* x_68; lean_object* x_69; -x_68 = lean_ctor_get(x_63, 1); -lean_inc(x_68); -lean_dec(x_63); -x_69 = lean_box(0); -x_3 = x_69; +lean_object* x_140; +lean_dec(x_135); +x_140 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_134); +return x_140; +} +} +else +{ +lean_object* x_141; lean_object* x_142; +x_141 = lean_ctor_get(x_131, 1); +lean_inc(x_141); +lean_dec(x_131); +x_142 = lean_box(0); +x_3 = x_142; x_8 = x_32; -x_10 = x_68; +x_10 = x_141; goto _start; } } else { -uint8_t x_71; +lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_71 = !lean_is_exclusive(x_63); -if (x_71 == 0) -{ -return x_63; +x_144 = lean_ctor_get(x_131, 0); +lean_inc(x_144); +x_145 = lean_ctor_get(x_131, 1); +lean_inc(x_145); +if (lean_is_exclusive(x_131)) { + lean_ctor_release(x_131, 0); + lean_ctor_release(x_131, 1); + x_146 = x_131; +} else { + lean_dec_ref(x_131); + x_146 = lean_box(0); } -else -{ -lean_object* x_72; lean_object* x_73; lean_object* x_74; -x_72 = lean_ctor_get(x_63, 0); -x_73 = lean_ctor_get(x_63, 1); -lean_inc(x_73); -lean_inc(x_72); -lean_dec(x_63); -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; +if (lean_is_scalar(x_146)) { + x_147 = lean_alloc_ctor(1, 2, 0); +} else { + x_147 = x_146; } +lean_ctor_set(x_147, 0, x_144); +lean_ctor_set(x_147, 1, x_145); +return x_147; } } else { -lean_object* x_75; lean_object* x_76; -x_75 = lean_ctor_get(x_59, 1); -lean_inc(x_75); -lean_dec(x_59); -x_76 = lean_box(0); -x_3 = x_76; +lean_object* x_148; lean_object* x_149; +x_148 = lean_ctor_get(x_127, 1); +lean_inc(x_148); +lean_dec(x_127); +x_149 = lean_box(0); +x_3 = x_149; x_8 = x_32; -x_10 = x_75; +x_10 = x_148; goto _start; } } else { -uint8_t x_78; +lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_78 = !lean_is_exclusive(x_59); -if (x_78 == 0) -{ -return x_59; +x_151 = lean_ctor_get(x_127, 0); +lean_inc(x_151); +x_152 = lean_ctor_get(x_127, 1); +lean_inc(x_152); +if (lean_is_exclusive(x_127)) { + lean_ctor_release(x_127, 0); + lean_ctor_release(x_127, 1); + x_153 = x_127; +} else { + lean_dec_ref(x_127); + x_153 = lean_box(0); } -else -{ -lean_object* x_79; lean_object* x_80; lean_object* x_81; -x_79 = lean_ctor_get(x_59, 0); -x_80 = lean_ctor_get(x_59, 1); -lean_inc(x_80); -lean_inc(x_79); -lean_dec(x_59); -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; +if (lean_is_scalar(x_153)) { + x_154 = lean_alloc_ctor(1, 2, 0); +} else { + x_154 = x_153; } +lean_ctor_set(x_154, 0, x_151); +lean_ctor_set(x_154, 1, x_152); +return x_154; } } else { -lean_object* x_82; lean_object* x_83; +lean_object* x_155; lean_object* x_156; lean_dec(x_42); -x_82 = lean_ctor_get(x_55, 1); -lean_inc(x_82); -lean_dec(x_55); -x_83 = lean_box(0); -x_3 = x_83; +x_155 = lean_ctor_get(x_123, 1); +lean_inc(x_155); +lean_dec(x_123); +x_156 = lean_box(0); +x_3 = x_156; x_8 = x_32; -x_10 = x_82; +x_10 = x_155; goto _start; } } else { -uint8_t x_85; +lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_dec(x_42); lean_dec(x_32); lean_dec(x_9); @@ -12405,43 +13308,45 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_85 = !lean_is_exclusive(x_55); -if (x_85 == 0) -{ -return x_55; +x_158 = lean_ctor_get(x_123, 0); +lean_inc(x_158); +x_159 = lean_ctor_get(x_123, 1); +lean_inc(x_159); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_160 = x_123; +} else { + lean_dec_ref(x_123); + x_160 = lean_box(0); } -else -{ -lean_object* x_86; lean_object* x_87; lean_object* x_88; -x_86 = lean_ctor_get(x_55, 0); -x_87 = lean_ctor_get(x_55, 1); -lean_inc(x_87); -lean_inc(x_86); -lean_dec(x_55); -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; +if (lean_is_scalar(x_160)) { + x_161 = lean_alloc_ctor(1, 2, 0); +} else { + x_161 = x_160; } +lean_ctor_set(x_161, 0, x_158); +lean_ctor_set(x_161, 1, x_159); +return x_161; } } else { -lean_object* x_89; lean_object* x_90; +lean_object* x_162; lean_object* x_163; lean_dec(x_42); -x_89 = lean_ctor_get(x_51, 1); -lean_inc(x_89); -lean_dec(x_51); -x_90 = lean_box(0); -x_3 = x_90; +x_162 = lean_ctor_get(x_119, 1); +lean_inc(x_162); +lean_dec(x_119); +x_163 = lean_box(0); +x_3 = x_163; x_8 = x_32; -x_10 = x_89; +x_10 = x_162; goto _start; } } else { -uint8_t x_92; +lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_dec(x_42); lean_dec(x_32); lean_dec(x_9); @@ -12449,29 +13354,31 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_92 = !lean_is_exclusive(x_51); -if (x_92 == 0) -{ -return x_51; +x_165 = lean_ctor_get(x_119, 0); +lean_inc(x_165); +x_166 = lean_ctor_get(x_119, 1); +lean_inc(x_166); +if (lean_is_exclusive(x_119)) { + lean_ctor_release(x_119, 0); + lean_ctor_release(x_119, 1); + x_167 = x_119; +} else { + lean_dec_ref(x_119); + x_167 = lean_box(0); } -else -{ -lean_object* x_93; lean_object* x_94; lean_object* x_95; -x_93 = lean_ctor_get(x_51, 0); -x_94 = lean_ctor_get(x_51, 1); -lean_inc(x_94); -lean_inc(x_93); -lean_dec(x_51); -x_95 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_95, 0, x_93); -lean_ctor_set(x_95, 1, x_94); -return x_95; +if (lean_is_scalar(x_167)) { + x_168 = lean_alloc_ctor(1, 2, 0); +} else { + x_168 = x_167; } +lean_ctor_set(x_168, 0, x_165); +lean_ctor_set(x_168, 1, x_166); +return x_168; } } else { -uint8_t x_96; +lean_object* x_169; lean_object* x_170; lean_dec(x_42); lean_dec(x_32); lean_dec(x_9); @@ -12479,47 +13386,31 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_96 = !lean_is_exclusive(x_43); -if (x_96 == 0) -{ -lean_object* x_97; lean_object* x_98; -x_97 = lean_ctor_get(x_43, 0); -lean_dec(x_97); -x_98 = lean_box(0); -lean_ctor_set(x_43, 0, x_98); -return x_43; -} -else -{ -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_43, 1); -lean_inc(x_99); -lean_dec(x_43); -x_100 = lean_box(0); -x_101 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_101, 0, x_100); -lean_ctor_set(x_101, 1, x_99); -return x_101; +x_169 = lean_box(0); +x_170 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_170, 0, x_169); +lean_ctor_set(x_170, 1, x_112); +return x_170; } } } else { -lean_object* x_102; lean_object* x_103; +lean_object* x_171; lean_object* x_172; lean_dec(x_42); -x_102 = lean_ctor_get(x_43, 1); -lean_inc(x_102); +x_171 = lean_ctor_get(x_43, 1); +lean_inc(x_171); lean_dec(x_43); -x_103 = lean_box(0); -x_3 = x_103; +x_172 = lean_box(0); +x_3 = x_172; x_8 = x_32; -x_10 = x_102; +x_10 = x_171; goto _start; } } else { -uint8_t x_105; +uint8_t x_174; lean_dec(x_42); lean_dec(x_32); lean_dec(x_9); @@ -12527,449 +13418,483 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_43); -if (x_105 == 0) +x_174 = !lean_is_exclusive(x_43); +if (x_174 == 0) { return x_43; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_43, 0); -x_107 = lean_ctor_get(x_43, 1); -lean_inc(x_107); -lean_inc(x_106); +lean_object* x_175; lean_object* x_176; lean_object* x_177; +x_175 = lean_ctor_get(x_43, 0); +x_176 = lean_ctor_get(x_43, 1); +lean_inc(x_176); +lean_inc(x_175); lean_dec(x_43); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +x_177 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_177, 0, x_175); +lean_ctor_set(x_177, 1, x_176); +return x_177; } } } else { -lean_object* x_109; +lean_object* x_178; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_109 = lean_box(0); -lean_ctor_set(x_33, 0, x_109); +x_178 = lean_box(0); +lean_ctor_set(x_33, 0, x_178); return x_33; } } else { -lean_object* x_110; lean_object* x_111; lean_object* x_112; uint8_t x_113; -x_110 = lean_ctor_get(x_33, 0); -x_111 = lean_ctor_get(x_33, 1); -lean_inc(x_111); -lean_inc(x_110); +lean_object* x_179; lean_object* x_180; lean_object* x_181; uint8_t x_182; +x_179 = lean_ctor_get(x_33, 0); +x_180 = lean_ctor_get(x_33, 1); +lean_inc(x_180); +lean_inc(x_179); lean_dec(x_33); -x_112 = lean_ctor_get(x_110, 2); -lean_inc(x_112); -lean_dec(x_110); -x_113 = l_List_isEmpty___rarg(x_112); -lean_dec(x_112); -if (x_113 == 0) +x_181 = lean_ctor_get(x_179, 2); +lean_inc(x_181); +lean_dec(x_179); +x_182 = l_List_isEmpty___rarg(x_181); +lean_dec(x_181); +if (x_182 == 0) { -uint8_t x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; -x_114 = 0; -x_115 = lean_box(x_114); -x_116 = lean_box(x_114); -x_117 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_117, 0, x_115); -lean_closure_set(x_117, 1, x_116); +uint8_t x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; +x_183 = 0; +x_184 = lean_box(x_183); +x_185 = lean_box(x_183); +x_186 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_186, 0, x_184); +lean_closure_set(x_186, 1, x_185); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_118 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_114, x_114, x_4, x_5, x_6, x_7, x_32, x_9, x_111); -if (lean_obj_tag(x_118) == 0) +x_187 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_183, x_183, x_4, x_5, x_6, x_7, x_32, x_9, x_180); +if (lean_obj_tag(x_187) == 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_188; uint8_t x_189; +x_188 = lean_ctor_get(x_187, 0); +lean_inc(x_188); +x_189 = lean_unbox(x_188); +lean_dec(x_188); +if (x_189 == 0) { -if (x_1 == 0) +lean_object* x_190; lean_object* x_191; uint8_t x_192; uint8_t x_193; +x_190 = lean_ctor_get(x_187, 1); +lean_inc(x_190); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_191 = x_187; +} else { + lean_dec_ref(x_187); + x_191 = lean_box(0); +} +x_192 = 0; +x_193 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_192); +if (x_193 == 0) { -lean_object* x_121; uint8_t x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; -x_121 = lean_ctor_get(x_118, 1); -lean_inc(x_121); -lean_dec(x_118); -x_122 = 1; -x_123 = lean_box(x_122); -x_124 = lean_box(x_114); -x_125 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_125, 0, x_123); -lean_closure_set(x_125, 1, x_124); +uint8_t x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; +lean_dec(x_191); +x_194 = 1; +x_195 = lean_box(x_194); +x_196 = lean_box(x_183); +x_197 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_197, 0, x_195); +lean_closure_set(x_197, 1, x_196); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_126 = l_Lean_Elab_Term_withoutPostponing___rarg(x_125, x_4, x_5, x_6, x_7, x_32, x_9, x_121); -if (lean_obj_tag(x_126) == 0) +x_198 = l_Lean_Elab_Term_withoutPostponing___rarg(x_197, x_4, x_5, x_6, x_7, x_32, x_9, x_190); +if (lean_obj_tag(x_198) == 0) { -lean_object* x_127; uint8_t x_128; -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -x_128 = lean_unbox(x_127); -lean_dec(x_127); -if (x_128 == 0) +lean_object* x_199; uint8_t x_200; +x_199 = lean_ctor_get(x_198, 0); +lean_inc(x_199); +x_200 = lean_unbox(x_199); +lean_dec(x_199); +if (x_200 == 0) { -lean_object* x_129; lean_object* x_130; -x_129 = lean_ctor_get(x_126, 1); -lean_inc(x_129); -lean_dec(x_126); +lean_object* x_201; lean_object* x_202; +x_201 = lean_ctor_get(x_198, 1); +lean_inc(x_201); +lean_dec(x_198); lean_inc(x_9); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_130 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_32, x_9, x_129); -if (lean_obj_tag(x_130) == 0) +x_202 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_32, x_9, x_201); +if (lean_obj_tag(x_202) == 0) { -lean_object* x_131; uint8_t x_132; -x_131 = lean_ctor_get(x_130, 0); -lean_inc(x_131); -x_132 = lean_unbox(x_131); -lean_dec(x_131); -if (x_132 == 0) +lean_object* x_203; uint8_t x_204; +x_203 = lean_ctor_get(x_202, 0); +lean_inc(x_203); +x_204 = lean_unbox(x_203); +lean_dec(x_203); +if (x_204 == 0) { -lean_object* x_133; lean_object* x_134; -x_133 = lean_ctor_get(x_130, 1); -lean_inc(x_133); -lean_dec(x_130); +lean_object* x_205; lean_object* x_206; +x_205 = lean_ctor_get(x_202, 1); +lean_inc(x_205); +lean_dec(x_202); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_134 = l_Lean_Elab_Term_withoutPostponing___rarg(x_117, x_4, x_5, x_6, x_7, x_32, x_9, x_133); -if (lean_obj_tag(x_134) == 0) +x_206 = l_Lean_Elab_Term_withoutPostponing___rarg(x_186, x_4, x_5, x_6, x_7, x_32, x_9, x_205); +if (lean_obj_tag(x_206) == 0) { -lean_object* x_135; uint8_t x_136; -x_135 = lean_ctor_get(x_134, 0); -lean_inc(x_135); -x_136 = lean_unbox(x_135); -lean_dec(x_135); -if (x_136 == 0) +lean_object* x_207; uint8_t x_208; +x_207 = lean_ctor_get(x_206, 0); +lean_inc(x_207); +x_208 = lean_unbox(x_207); +lean_dec(x_207); +if (x_208 == 0) { -lean_object* x_137; lean_object* x_138; -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -lean_dec(x_134); +lean_object* x_209; lean_object* x_210; +x_209 = lean_ctor_get(x_206, 1); +lean_inc(x_209); +lean_dec(x_206); lean_inc(x_9); lean_inc(x_32); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_138 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_114, x_122, x_4, x_5, x_6, x_7, x_32, x_9, x_137); -if (lean_obj_tag(x_138) == 0) +x_210 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_183, x_194, x_4, x_5, x_6, x_7, x_32, x_9, x_209); +if (lean_obj_tag(x_210) == 0) { -lean_object* x_139; uint8_t x_140; -x_139 = lean_ctor_get(x_138, 0); -lean_inc(x_139); -x_140 = lean_unbox(x_139); -lean_dec(x_139); -if (x_140 == 0) +lean_object* x_211; uint8_t x_212; +x_211 = lean_ctor_get(x_210, 0); +lean_inc(x_211); +x_212 = lean_unbox(x_211); +lean_dec(x_211); +if (x_212 == 0) +{ +lean_object* x_213; lean_object* x_214; uint8_t x_215; uint8_t x_216; +x_213 = lean_ctor_get(x_210, 1); +lean_inc(x_213); +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + x_214 = x_210; +} else { + lean_dec_ref(x_210); + x_214 = lean_box(0); +} +x_215 = 1; +x_216 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_215); +if (x_216 == 0) { -lean_object* x_141; lean_object* x_142; -x_141 = lean_ctor_get(x_138, 1); -lean_inc(x_141); -lean_dec(x_138); -x_142 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_141); -return x_142; +lean_object* x_217; lean_object* x_218; +lean_dec(x_32); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_217 = lean_box(0); +if (lean_is_scalar(x_214)) { + x_218 = lean_alloc_ctor(0, 2, 0); +} else { + x_218 = x_214; +} +lean_ctor_set(x_218, 0, x_217); +lean_ctor_set(x_218, 1, x_213); +return x_218; } else { -lean_object* x_143; lean_object* x_144; -x_143 = lean_ctor_get(x_138, 1); -lean_inc(x_143); -lean_dec(x_138); -x_144 = lean_box(0); -x_3 = x_144; +lean_object* x_219; +lean_dec(x_214); +x_219 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_32, x_9, x_213); +return x_219; +} +} +else +{ +lean_object* x_220; lean_object* x_221; +x_220 = lean_ctor_get(x_210, 1); +lean_inc(x_220); +lean_dec(x_210); +x_221 = lean_box(0); +x_3 = x_221; x_8 = x_32; -x_10 = x_143; +x_10 = x_220; goto _start; } } else { -lean_object* x_146; lean_object* x_147; lean_object* x_148; lean_object* x_149; +lean_object* x_223; lean_object* x_224; lean_object* x_225; lean_object* x_226; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_146 = lean_ctor_get(x_138, 0); -lean_inc(x_146); -x_147 = lean_ctor_get(x_138, 1); -lean_inc(x_147); -if (lean_is_exclusive(x_138)) { - lean_ctor_release(x_138, 0); - lean_ctor_release(x_138, 1); - x_148 = x_138; +x_223 = lean_ctor_get(x_210, 0); +lean_inc(x_223); +x_224 = lean_ctor_get(x_210, 1); +lean_inc(x_224); +if (lean_is_exclusive(x_210)) { + lean_ctor_release(x_210, 0); + lean_ctor_release(x_210, 1); + x_225 = x_210; } else { - lean_dec_ref(x_138); - x_148 = lean_box(0); + lean_dec_ref(x_210); + x_225 = lean_box(0); } -if (lean_is_scalar(x_148)) { - x_149 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_225)) { + x_226 = lean_alloc_ctor(1, 2, 0); } else { - x_149 = x_148; + x_226 = x_225; } -lean_ctor_set(x_149, 0, x_146); -lean_ctor_set(x_149, 1, x_147); -return x_149; +lean_ctor_set(x_226, 0, x_223); +lean_ctor_set(x_226, 1, x_224); +return x_226; } } else { -lean_object* x_150; lean_object* x_151; -x_150 = lean_ctor_get(x_134, 1); -lean_inc(x_150); -lean_dec(x_134); -x_151 = lean_box(0); -x_3 = x_151; +lean_object* x_227; lean_object* x_228; +x_227 = lean_ctor_get(x_206, 1); +lean_inc(x_227); +lean_dec(x_206); +x_228 = lean_box(0); +x_3 = x_228; x_8 = x_32; -x_10 = x_150; +x_10 = x_227; goto _start; } } else { -lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; +lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_153 = lean_ctor_get(x_134, 0); -lean_inc(x_153); -x_154 = lean_ctor_get(x_134, 1); -lean_inc(x_154); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - x_155 = x_134; +x_230 = lean_ctor_get(x_206, 0); +lean_inc(x_230); +x_231 = lean_ctor_get(x_206, 1); +lean_inc(x_231); +if (lean_is_exclusive(x_206)) { + lean_ctor_release(x_206, 0); + lean_ctor_release(x_206, 1); + x_232 = x_206; } else { - lean_dec_ref(x_134); - x_155 = lean_box(0); + lean_dec_ref(x_206); + x_232 = lean_box(0); } -if (lean_is_scalar(x_155)) { - x_156 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_232)) { + x_233 = lean_alloc_ctor(1, 2, 0); } else { - x_156 = x_155; + x_233 = x_232; } -lean_ctor_set(x_156, 0, x_153); -lean_ctor_set(x_156, 1, x_154); -return x_156; +lean_ctor_set(x_233, 0, x_230); +lean_ctor_set(x_233, 1, x_231); +return x_233; } } else { -lean_object* x_157; lean_object* x_158; -lean_dec(x_117); -x_157 = lean_ctor_get(x_130, 1); -lean_inc(x_157); -lean_dec(x_130); -x_158 = lean_box(0); -x_3 = x_158; +lean_object* x_234; lean_object* x_235; +lean_dec(x_186); +x_234 = lean_ctor_get(x_202, 1); +lean_inc(x_234); +lean_dec(x_202); +x_235 = lean_box(0); +x_3 = x_235; x_8 = x_32; -x_10 = x_157; +x_10 = x_234; goto _start; } } else { -lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -lean_dec(x_117); +lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; +lean_dec(x_186); lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_160 = lean_ctor_get(x_130, 0); -lean_inc(x_160); -x_161 = lean_ctor_get(x_130, 1); -lean_inc(x_161); -if (lean_is_exclusive(x_130)) { - lean_ctor_release(x_130, 0); - lean_ctor_release(x_130, 1); - x_162 = x_130; +x_237 = lean_ctor_get(x_202, 0); +lean_inc(x_237); +x_238 = lean_ctor_get(x_202, 1); +lean_inc(x_238); +if (lean_is_exclusive(x_202)) { + lean_ctor_release(x_202, 0); + lean_ctor_release(x_202, 1); + x_239 = x_202; } else { - lean_dec_ref(x_130); - x_162 = lean_box(0); + lean_dec_ref(x_202); + x_239 = lean_box(0); } -if (lean_is_scalar(x_162)) { - x_163 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_239)) { + x_240 = lean_alloc_ctor(1, 2, 0); } else { - x_163 = x_162; + x_240 = x_239; } -lean_ctor_set(x_163, 0, x_160); -lean_ctor_set(x_163, 1, x_161); -return x_163; +lean_ctor_set(x_240, 0, x_237); +lean_ctor_set(x_240, 1, x_238); +return x_240; } } else { -lean_object* x_164; lean_object* x_165; -lean_dec(x_117); -x_164 = lean_ctor_get(x_126, 1); -lean_inc(x_164); -lean_dec(x_126); -x_165 = lean_box(0); -x_3 = x_165; +lean_object* x_241; lean_object* x_242; +lean_dec(x_186); +x_241 = lean_ctor_get(x_198, 1); +lean_inc(x_241); +lean_dec(x_198); +x_242 = lean_box(0); +x_3 = x_242; x_8 = x_32; -x_10 = x_164; +x_10 = x_241; goto _start; } } else { -lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -lean_dec(x_117); +lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; +lean_dec(x_186); lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_167 = lean_ctor_get(x_126, 0); -lean_inc(x_167); -x_168 = lean_ctor_get(x_126, 1); -lean_inc(x_168); -if (lean_is_exclusive(x_126)) { - lean_ctor_release(x_126, 0); - lean_ctor_release(x_126, 1); - x_169 = x_126; +x_244 = lean_ctor_get(x_198, 0); +lean_inc(x_244); +x_245 = lean_ctor_get(x_198, 1); +lean_inc(x_245); +if (lean_is_exclusive(x_198)) { + lean_ctor_release(x_198, 0); + lean_ctor_release(x_198, 1); + x_246 = x_198; } else { - lean_dec_ref(x_126); - x_169 = lean_box(0); + lean_dec_ref(x_198); + x_246 = lean_box(0); } -if (lean_is_scalar(x_169)) { - x_170 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_246)) { + x_247 = lean_alloc_ctor(1, 2, 0); } else { - x_170 = x_169; + x_247 = x_246; } -lean_ctor_set(x_170, 0, x_167); -lean_ctor_set(x_170, 1, x_168); -return x_170; +lean_ctor_set(x_247, 0, x_244); +lean_ctor_set(x_247, 1, x_245); +return x_247; } } else { -lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; -lean_dec(x_117); +lean_object* x_248; lean_object* x_249; +lean_dec(x_186); lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_171 = lean_ctor_get(x_118, 1); -lean_inc(x_171); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_172 = x_118; -} else { - lean_dec_ref(x_118); - x_172 = lean_box(0); -} -x_173 = lean_box(0); -if (lean_is_scalar(x_172)) { - x_174 = lean_alloc_ctor(0, 2, 0); +x_248 = lean_box(0); +if (lean_is_scalar(x_191)) { + x_249 = lean_alloc_ctor(0, 2, 0); } else { - x_174 = x_172; + x_249 = x_191; } -lean_ctor_set(x_174, 0, x_173); -lean_ctor_set(x_174, 1, x_171); -return x_174; +lean_ctor_set(x_249, 0, x_248); +lean_ctor_set(x_249, 1, x_190); +return x_249; } } else { -lean_object* x_175; lean_object* x_176; -lean_dec(x_117); -x_175 = lean_ctor_get(x_118, 1); -lean_inc(x_175); -lean_dec(x_118); -x_176 = lean_box(0); -x_3 = x_176; +lean_object* x_250; lean_object* x_251; +lean_dec(x_186); +x_250 = lean_ctor_get(x_187, 1); +lean_inc(x_250); +lean_dec(x_187); +x_251 = lean_box(0); +x_3 = x_251; x_8 = x_32; -x_10 = x_175; +x_10 = x_250; goto _start; } } else { -lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -lean_dec(x_117); +lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +lean_dec(x_186); lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_178 = lean_ctor_get(x_118, 0); -lean_inc(x_178); -x_179 = lean_ctor_get(x_118, 1); -lean_inc(x_179); -if (lean_is_exclusive(x_118)) { - lean_ctor_release(x_118, 0); - lean_ctor_release(x_118, 1); - x_180 = x_118; +x_253 = lean_ctor_get(x_187, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_187, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_187)) { + lean_ctor_release(x_187, 0); + lean_ctor_release(x_187, 1); + x_255 = x_187; } else { - lean_dec_ref(x_118); - x_180 = lean_box(0); + lean_dec_ref(x_187); + x_255 = lean_box(0); } -if (lean_is_scalar(x_180)) { - x_181 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(1, 2, 0); } else { - x_181 = x_180; + x_256 = x_255; } -lean_ctor_set(x_181, 0, x_178); -lean_ctor_set(x_181, 1, x_179); -return x_181; +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_254); +return x_256; } } else { -lean_object* x_182; lean_object* x_183; +lean_object* x_257; lean_object* x_258; lean_dec(x_32); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_182 = lean_box(0); -x_183 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_183, 0, x_182); -lean_ctor_set(x_183, 1, x_111); -return x_183; +x_257 = lean_box(0); +x_258 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_258, 0, x_257); +lean_ctor_set(x_258, 1, x_180); +return x_258; } } } else { -lean_object* x_184; +lean_object* x_259; lean_dec(x_25); lean_dec(x_24); lean_dec(x_23); @@ -12980,559 +13905,593 @@ lean_dec(x_18); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); -x_184 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_13); +x_259 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(x_28, x_4, x_5, x_6, x_7, x_8, x_9, x_13); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_184; -} -} -else -{ -lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; lean_object* x_189; lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; uint8_t x_196; uint8_t x_197; lean_object* x_198; lean_object* x_199; uint8_t x_200; -x_185 = lean_ctor_get(x_8, 0); -x_186 = lean_ctor_get(x_8, 1); -x_187 = lean_ctor_get(x_8, 2); -x_188 = lean_ctor_get(x_8, 3); -x_189 = lean_ctor_get(x_8, 4); -x_190 = lean_ctor_get(x_8, 5); -x_191 = lean_ctor_get(x_8, 6); -x_192 = lean_ctor_get(x_8, 7); -x_193 = lean_ctor_get(x_8, 8); -x_194 = lean_ctor_get(x_8, 9); -x_195 = lean_ctor_get(x_8, 10); -x_196 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); -x_197 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); -lean_inc(x_195); -lean_inc(x_194); -lean_inc(x_193); -lean_inc(x_192); -lean_inc(x_191); -lean_inc(x_190); -lean_inc(x_189); -lean_inc(x_188); -lean_inc(x_187); -lean_inc(x_186); -lean_inc(x_185); +return x_259; +} +} +else +{ +lean_object* x_260; lean_object* x_261; lean_object* x_262; lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; uint8_t x_271; uint8_t x_272; lean_object* x_273; lean_object* x_274; uint8_t x_275; +x_260 = lean_ctor_get(x_8, 0); +x_261 = lean_ctor_get(x_8, 1); +x_262 = lean_ctor_get(x_8, 2); +x_263 = lean_ctor_get(x_8, 3); +x_264 = lean_ctor_get(x_8, 4); +x_265 = lean_ctor_get(x_8, 5); +x_266 = lean_ctor_get(x_8, 6); +x_267 = lean_ctor_get(x_8, 7); +x_268 = lean_ctor_get(x_8, 8); +x_269 = lean_ctor_get(x_8, 9); +x_270 = lean_ctor_get(x_8, 10); +x_271 = lean_ctor_get_uint8(x_8, sizeof(void*)*11); +x_272 = lean_ctor_get_uint8(x_8, sizeof(void*)*11 + 1); +lean_inc(x_270); +lean_inc(x_269); +lean_inc(x_268); +lean_inc(x_267); +lean_inc(x_266); +lean_inc(x_265); +lean_inc(x_264); +lean_inc(x_263); +lean_inc(x_262); +lean_inc(x_261); +lean_inc(x_260); lean_dec(x_8); -x_198 = l_Lean_replaceRef(x_12, x_190); -lean_dec(x_190); +x_273 = l_Lean_replaceRef(x_12, x_265); +lean_dec(x_265); lean_dec(x_12); -lean_inc(x_195); -lean_inc(x_194); -lean_inc(x_193); -lean_inc(x_192); -lean_inc(x_191); -lean_inc(x_198); -lean_inc(x_189); -lean_inc(x_188); -lean_inc(x_187); -lean_inc(x_186); -lean_inc(x_185); -x_199 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_199, 0, x_185); -lean_ctor_set(x_199, 1, x_186); -lean_ctor_set(x_199, 2, x_187); -lean_ctor_set(x_199, 3, x_188); -lean_ctor_set(x_199, 4, x_189); -lean_ctor_set(x_199, 5, x_198); -lean_ctor_set(x_199, 6, x_191); -lean_ctor_set(x_199, 7, x_192); -lean_ctor_set(x_199, 8, x_193); -lean_ctor_set(x_199, 9, x_194); -lean_ctor_set(x_199, 10, x_195); -lean_ctor_set_uint8(x_199, sizeof(void*)*11, x_196); -lean_ctor_set_uint8(x_199, sizeof(void*)*11 + 1, x_197); -x_200 = lean_nat_dec_eq(x_188, x_189); -if (x_200 == 0) -{ -lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; uint8_t x_209; -lean_dec(x_199); -x_201 = lean_unsigned_to_nat(1u); -x_202 = lean_nat_add(x_188, x_201); -lean_dec(x_188); -x_203 = lean_alloc_ctor(0, 11, 2); -lean_ctor_set(x_203, 0, x_185); -lean_ctor_set(x_203, 1, x_186); -lean_ctor_set(x_203, 2, x_187); -lean_ctor_set(x_203, 3, x_202); -lean_ctor_set(x_203, 4, x_189); -lean_ctor_set(x_203, 5, x_198); -lean_ctor_set(x_203, 6, x_191); -lean_ctor_set(x_203, 7, x_192); -lean_ctor_set(x_203, 8, x_193); -lean_ctor_set(x_203, 9, x_194); -lean_ctor_set(x_203, 10, x_195); -lean_ctor_set_uint8(x_203, sizeof(void*)*11, x_196); -lean_ctor_set_uint8(x_203, sizeof(void*)*11 + 1, x_197); -x_204 = lean_st_ref_get(x_5, x_13); -x_205 = lean_ctor_get(x_204, 0); -lean_inc(x_205); -x_206 = lean_ctor_get(x_204, 1); -lean_inc(x_206); -if (lean_is_exclusive(x_204)) { - lean_ctor_release(x_204, 0); - lean_ctor_release(x_204, 1); - x_207 = x_204; +lean_inc(x_270); +lean_inc(x_269); +lean_inc(x_268); +lean_inc(x_267); +lean_inc(x_266); +lean_inc(x_273); +lean_inc(x_264); +lean_inc(x_263); +lean_inc(x_262); +lean_inc(x_261); +lean_inc(x_260); +x_274 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_274, 0, x_260); +lean_ctor_set(x_274, 1, x_261); +lean_ctor_set(x_274, 2, x_262); +lean_ctor_set(x_274, 3, x_263); +lean_ctor_set(x_274, 4, x_264); +lean_ctor_set(x_274, 5, x_273); +lean_ctor_set(x_274, 6, x_266); +lean_ctor_set(x_274, 7, x_267); +lean_ctor_set(x_274, 8, x_268); +lean_ctor_set(x_274, 9, x_269); +lean_ctor_set(x_274, 10, x_270); +lean_ctor_set_uint8(x_274, sizeof(void*)*11, x_271); +lean_ctor_set_uint8(x_274, sizeof(void*)*11 + 1, x_272); +x_275 = lean_nat_dec_eq(x_263, x_264); +if (x_275 == 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_object* x_283; uint8_t x_284; +lean_dec(x_274); +x_276 = lean_unsigned_to_nat(1u); +x_277 = lean_nat_add(x_263, x_276); +lean_dec(x_263); +x_278 = lean_alloc_ctor(0, 11, 2); +lean_ctor_set(x_278, 0, x_260); +lean_ctor_set(x_278, 1, x_261); +lean_ctor_set(x_278, 2, x_262); +lean_ctor_set(x_278, 3, x_277); +lean_ctor_set(x_278, 4, x_264); +lean_ctor_set(x_278, 5, x_273); +lean_ctor_set(x_278, 6, x_266); +lean_ctor_set(x_278, 7, x_267); +lean_ctor_set(x_278, 8, x_268); +lean_ctor_set(x_278, 9, x_269); +lean_ctor_set(x_278, 10, x_270); +lean_ctor_set_uint8(x_278, sizeof(void*)*11, x_271); +lean_ctor_set_uint8(x_278, sizeof(void*)*11 + 1, x_272); +x_279 = lean_st_ref_get(x_5, x_13); +x_280 = lean_ctor_get(x_279, 0); +lean_inc(x_280); +x_281 = lean_ctor_get(x_279, 1); +lean_inc(x_281); +if (lean_is_exclusive(x_279)) { + lean_ctor_release(x_279, 0); + lean_ctor_release(x_279, 1); + x_282 = x_279; } else { - lean_dec_ref(x_204); - x_207 = lean_box(0); + lean_dec_ref(x_279); + x_282 = lean_box(0); } -x_208 = lean_ctor_get(x_205, 2); -lean_inc(x_208); -lean_dec(x_205); -x_209 = l_List_isEmpty___rarg(x_208); -lean_dec(x_208); -if (x_209 == 0) -{ -uint8_t x_210; lean_object* x_211; lean_object* x_212; lean_object* x_213; lean_object* x_214; -lean_dec(x_207); -x_210 = 0; -x_211 = lean_box(x_210); -x_212 = lean_box(x_210); -x_213 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_213, 0, x_211); -lean_closure_set(x_213, 1, x_212); +x_283 = lean_ctor_get(x_280, 2); +lean_inc(x_283); +lean_dec(x_280); +x_284 = l_List_isEmpty___rarg(x_283); +lean_dec(x_283); +if (x_284 == 0) +{ +uint8_t x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; +lean_dec(x_282); +x_285 = 0; +x_286 = lean_box(x_285); +x_287 = lean_box(x_285); +x_288 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_288, 0, x_286); +lean_closure_set(x_288, 1, x_287); lean_inc(x_9); -lean_inc(x_203); +lean_inc(x_278); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_214 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_210, x_210, x_4, x_5, x_6, x_7, x_203, x_9, x_206); -if (lean_obj_tag(x_214) == 0) -{ -lean_object* x_215; uint8_t x_216; -x_215 = lean_ctor_get(x_214, 0); -lean_inc(x_215); -x_216 = lean_unbox(x_215); -lean_dec(x_215); -if (x_216 == 0) -{ -if (x_1 == 0) -{ -lean_object* x_217; uint8_t x_218; lean_object* x_219; lean_object* x_220; lean_object* x_221; lean_object* x_222; -x_217 = lean_ctor_get(x_214, 1); -lean_inc(x_217); -lean_dec(x_214); -x_218 = 1; -x_219 = lean_box(x_218); -x_220 = lean_box(x_210); -x_221 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); -lean_closure_set(x_221, 0, x_219); -lean_closure_set(x_221, 1, x_220); +x_289 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_285, x_285, x_4, x_5, x_6, x_7, x_278, x_9, x_281); +if (lean_obj_tag(x_289) == 0) +{ +lean_object* x_290; uint8_t x_291; +x_290 = lean_ctor_get(x_289, 0); +lean_inc(x_290); +x_291 = lean_unbox(x_290); +lean_dec(x_290); +if (x_291 == 0) +{ +lean_object* x_292; lean_object* x_293; uint8_t x_294; uint8_t x_295; +x_292 = lean_ctor_get(x_289, 1); +lean_inc(x_292); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + x_293 = x_289; +} else { + lean_dec_ref(x_289); + x_293 = lean_box(0); +} +x_294 = 0; +x_295 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_294); +if (x_295 == 0) +{ +uint8_t x_296; lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; +lean_dec(x_293); +x_296 = 1; +x_297 = lean_box(x_296); +x_298 = lean_box(x_285); +x_299 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep___boxed), 9, 2); +lean_closure_set(x_299, 0, x_297); +lean_closure_set(x_299, 1, x_298); lean_inc(x_9); -lean_inc(x_203); +lean_inc(x_278); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_222 = l_Lean_Elab_Term_withoutPostponing___rarg(x_221, x_4, x_5, x_6, x_7, x_203, x_9, x_217); -if (lean_obj_tag(x_222) == 0) -{ -lean_object* x_223; uint8_t x_224; -x_223 = lean_ctor_get(x_222, 0); -lean_inc(x_223); -x_224 = lean_unbox(x_223); -lean_dec(x_223); -if (x_224 == 0) -{ -lean_object* x_225; lean_object* x_226; -x_225 = lean_ctor_get(x_222, 1); -lean_inc(x_225); -lean_dec(x_222); +x_300 = l_Lean_Elab_Term_withoutPostponing___rarg(x_299, x_4, x_5, x_6, x_7, x_278, x_9, x_292); +if (lean_obj_tag(x_300) == 0) +{ +lean_object* x_301; uint8_t x_302; +x_301 = lean_ctor_get(x_300, 0); +lean_inc(x_301); +x_302 = lean_unbox(x_301); +lean_dec(x_301); +if (x_302 == 0) +{ +lean_object* x_303; lean_object* x_304; +x_303 = lean_ctor_get(x_300, 1); +lean_inc(x_303); +lean_dec(x_300); lean_inc(x_9); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_226 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_203, x_9, x_225); -if (lean_obj_tag(x_226) == 0) -{ -lean_object* x_227; uint8_t x_228; -x_227 = lean_ctor_get(x_226, 0); -lean_inc(x_227); -x_228 = lean_unbox(x_227); -lean_dec(x_227); -if (x_228 == 0) -{ -lean_object* x_229; lean_object* x_230; -x_229 = lean_ctor_get(x_226, 1); -lean_inc(x_229); -lean_dec(x_226); +x_304 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefault(x_4, x_5, x_6, x_7, x_278, x_9, x_303); +if (lean_obj_tag(x_304) == 0) +{ +lean_object* x_305; uint8_t x_306; +x_305 = lean_ctor_get(x_304, 0); +lean_inc(x_305); +x_306 = lean_unbox(x_305); +lean_dec(x_305); +if (x_306 == 0) +{ +lean_object* x_307; lean_object* x_308; +x_307 = lean_ctor_get(x_304, 1); +lean_inc(x_307); +lean_dec(x_304); lean_inc(x_9); -lean_inc(x_203); +lean_inc(x_278); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_230 = l_Lean_Elab_Term_withoutPostponing___rarg(x_213, x_4, x_5, x_6, x_7, x_203, x_9, x_229); -if (lean_obj_tag(x_230) == 0) -{ -lean_object* x_231; uint8_t x_232; -x_231 = lean_ctor_get(x_230, 0); -lean_inc(x_231); -x_232 = lean_unbox(x_231); -lean_dec(x_231); -if (x_232 == 0) -{ -lean_object* x_233; lean_object* x_234; -x_233 = lean_ctor_get(x_230, 1); -lean_inc(x_233); -lean_dec(x_230); +x_308 = l_Lean_Elab_Term_withoutPostponing___rarg(x_288, x_4, x_5, x_6, x_7, x_278, x_9, x_307); +if (lean_obj_tag(x_308) == 0) +{ +lean_object* x_309; uint8_t x_310; +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_unbox(x_309); +lean_dec(x_309); +if (x_310 == 0) +{ +lean_object* x_311; lean_object* x_312; +x_311 = lean_ctor_get(x_308, 1); +lean_inc(x_311); +lean_dec(x_308); lean_inc(x_9); -lean_inc(x_203); +lean_inc(x_278); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_234 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_210, x_218, x_4, x_5, x_6, x_7, x_203, x_9, x_233); -if (lean_obj_tag(x_234) == 0) +x_312 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeSyntheticMVarsStep(x_285, x_296, x_4, x_5, x_6, x_7, x_278, x_9, x_311); +if (lean_obj_tag(x_312) == 0) +{ +lean_object* x_313; uint8_t x_314; +x_313 = lean_ctor_get(x_312, 0); +lean_inc(x_313); +x_314 = lean_unbox(x_313); +lean_dec(x_313); +if (x_314 == 0) +{ +lean_object* x_315; lean_object* x_316; uint8_t x_317; uint8_t x_318; +x_315 = lean_ctor_get(x_312, 1); +lean_inc(x_315); +if (lean_is_exclusive(x_312)) { + lean_ctor_release(x_312, 0); + lean_ctor_release(x_312, 1); + x_316 = x_312; +} else { + lean_dec_ref(x_312); + x_316 = lean_box(0); +} +x_317 = 1; +x_318 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_317); +if (x_318 == 0) { -lean_object* x_235; uint8_t x_236; -x_235 = lean_ctor_get(x_234, 0); -lean_inc(x_235); -x_236 = lean_unbox(x_235); -lean_dec(x_235); -if (x_236 == 0) +lean_object* x_319; lean_object* x_320; +lean_dec(x_278); +lean_dec(x_9); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +x_319 = lean_box(0); +if (lean_is_scalar(x_316)) { + x_320 = lean_alloc_ctor(0, 2, 0); +} else { + x_320 = x_316; +} +lean_ctor_set(x_320, 0, x_319); +lean_ctor_set(x_320, 1, x_315); +return x_320; +} +else { -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_234, 1); -lean_inc(x_237); -lean_dec(x_234); -x_238 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_203, x_9, x_237); -return x_238; +lean_object* x_321; +lean_dec(x_316); +x_321 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reportStuckSyntheticMVars(x_2, x_4, x_5, x_6, x_7, x_278, x_9, x_315); +return x_321; +} } else { -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_234, 1); -lean_inc(x_239); -lean_dec(x_234); -x_240 = lean_box(0); -x_3 = x_240; -x_8 = x_203; -x_10 = x_239; +lean_object* x_322; lean_object* x_323; +x_322 = lean_ctor_get(x_312, 1); +lean_inc(x_322); +lean_dec(x_312); +x_323 = lean_box(0); +x_3 = x_323; +x_8 = x_278; +x_10 = x_322; goto _start; } } else { -lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; -lean_dec(x_203); +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_242 = lean_ctor_get(x_234, 0); -lean_inc(x_242); -x_243 = lean_ctor_get(x_234, 1); -lean_inc(x_243); -if (lean_is_exclusive(x_234)) { - lean_ctor_release(x_234, 0); - lean_ctor_release(x_234, 1); - x_244 = x_234; +x_325 = lean_ctor_get(x_312, 0); +lean_inc(x_325); +x_326 = lean_ctor_get(x_312, 1); +lean_inc(x_326); +if (lean_is_exclusive(x_312)) { + lean_ctor_release(x_312, 0); + lean_ctor_release(x_312, 1); + x_327 = x_312; } else { - lean_dec_ref(x_234); - x_244 = lean_box(0); + lean_dec_ref(x_312); + x_327 = lean_box(0); } -if (lean_is_scalar(x_244)) { - x_245 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_327)) { + x_328 = lean_alloc_ctor(1, 2, 0); } else { - x_245 = x_244; + x_328 = x_327; } -lean_ctor_set(x_245, 0, x_242); -lean_ctor_set(x_245, 1, x_243); -return x_245; +lean_ctor_set(x_328, 0, x_325); +lean_ctor_set(x_328, 1, x_326); +return x_328; } } else { -lean_object* x_246; lean_object* x_247; -x_246 = lean_ctor_get(x_230, 1); -lean_inc(x_246); -lean_dec(x_230); -x_247 = lean_box(0); -x_3 = x_247; -x_8 = x_203; -x_10 = x_246; +lean_object* x_329; lean_object* x_330; +x_329 = lean_ctor_get(x_308, 1); +lean_inc(x_329); +lean_dec(x_308); +x_330 = lean_box(0); +x_3 = x_330; +x_8 = x_278; +x_10 = x_329; goto _start; } } else { -lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; -lean_dec(x_203); +lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_249 = lean_ctor_get(x_230, 0); -lean_inc(x_249); -x_250 = lean_ctor_get(x_230, 1); -lean_inc(x_250); -if (lean_is_exclusive(x_230)) { - lean_ctor_release(x_230, 0); - lean_ctor_release(x_230, 1); - x_251 = x_230; +x_332 = lean_ctor_get(x_308, 0); +lean_inc(x_332); +x_333 = lean_ctor_get(x_308, 1); +lean_inc(x_333); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_334 = x_308; } else { - lean_dec_ref(x_230); - x_251 = lean_box(0); + lean_dec_ref(x_308); + x_334 = lean_box(0); } -if (lean_is_scalar(x_251)) { - x_252 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_334)) { + x_335 = lean_alloc_ctor(1, 2, 0); } else { - x_252 = x_251; + x_335 = x_334; } -lean_ctor_set(x_252, 0, x_249); -lean_ctor_set(x_252, 1, x_250); -return x_252; +lean_ctor_set(x_335, 0, x_332); +lean_ctor_set(x_335, 1, x_333); +return x_335; } } else { -lean_object* x_253; lean_object* x_254; -lean_dec(x_213); -x_253 = lean_ctor_get(x_226, 1); -lean_inc(x_253); -lean_dec(x_226); -x_254 = lean_box(0); -x_3 = x_254; -x_8 = x_203; -x_10 = x_253; +lean_object* x_336; lean_object* x_337; +lean_dec(x_288); +x_336 = lean_ctor_get(x_304, 1); +lean_inc(x_336); +lean_dec(x_304); +x_337 = lean_box(0); +x_3 = x_337; +x_8 = x_278; +x_10 = x_336; goto _start; } } else { -lean_object* x_256; lean_object* x_257; lean_object* x_258; lean_object* x_259; -lean_dec(x_213); -lean_dec(x_203); +lean_object* x_339; lean_object* x_340; lean_object* x_341; lean_object* x_342; +lean_dec(x_288); +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_256 = lean_ctor_get(x_226, 0); -lean_inc(x_256); -x_257 = lean_ctor_get(x_226, 1); -lean_inc(x_257); -if (lean_is_exclusive(x_226)) { - lean_ctor_release(x_226, 0); - lean_ctor_release(x_226, 1); - x_258 = x_226; +x_339 = lean_ctor_get(x_304, 0); +lean_inc(x_339); +x_340 = lean_ctor_get(x_304, 1); +lean_inc(x_340); +if (lean_is_exclusive(x_304)) { + lean_ctor_release(x_304, 0); + lean_ctor_release(x_304, 1); + x_341 = x_304; } else { - lean_dec_ref(x_226); - x_258 = lean_box(0); + lean_dec_ref(x_304); + x_341 = lean_box(0); } -if (lean_is_scalar(x_258)) { - x_259 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_341)) { + x_342 = lean_alloc_ctor(1, 2, 0); } else { - x_259 = x_258; + x_342 = x_341; } -lean_ctor_set(x_259, 0, x_256); -lean_ctor_set(x_259, 1, x_257); -return x_259; +lean_ctor_set(x_342, 0, x_339); +lean_ctor_set(x_342, 1, x_340); +return x_342; } } else { -lean_object* x_260; lean_object* x_261; -lean_dec(x_213); -x_260 = lean_ctor_get(x_222, 1); -lean_inc(x_260); -lean_dec(x_222); -x_261 = lean_box(0); -x_3 = x_261; -x_8 = x_203; -x_10 = x_260; +lean_object* x_343; lean_object* x_344; +lean_dec(x_288); +x_343 = lean_ctor_get(x_300, 1); +lean_inc(x_343); +lean_dec(x_300); +x_344 = lean_box(0); +x_3 = x_344; +x_8 = x_278; +x_10 = x_343; goto _start; } } else { -lean_object* x_263; lean_object* x_264; lean_object* x_265; lean_object* x_266; -lean_dec(x_213); -lean_dec(x_203); +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; +lean_dec(x_288); +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_263 = lean_ctor_get(x_222, 0); -lean_inc(x_263); -x_264 = lean_ctor_get(x_222, 1); -lean_inc(x_264); -if (lean_is_exclusive(x_222)) { - lean_ctor_release(x_222, 0); - lean_ctor_release(x_222, 1); - x_265 = x_222; +x_346 = lean_ctor_get(x_300, 0); +lean_inc(x_346); +x_347 = lean_ctor_get(x_300, 1); +lean_inc(x_347); +if (lean_is_exclusive(x_300)) { + lean_ctor_release(x_300, 0); + lean_ctor_release(x_300, 1); + x_348 = x_300; } else { - lean_dec_ref(x_222); - x_265 = lean_box(0); + lean_dec_ref(x_300); + x_348 = lean_box(0); } -if (lean_is_scalar(x_265)) { - x_266 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_348)) { + x_349 = lean_alloc_ctor(1, 2, 0); } else { - x_266 = x_265; + x_349 = x_348; } -lean_ctor_set(x_266, 0, x_263); -lean_ctor_set(x_266, 1, x_264); -return x_266; +lean_ctor_set(x_349, 0, x_346); +lean_ctor_set(x_349, 1, x_347); +return x_349; } } else { -lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; -lean_dec(x_213); -lean_dec(x_203); +lean_object* x_350; lean_object* x_351; +lean_dec(x_288); +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_267 = lean_ctor_get(x_214, 1); -lean_inc(x_267); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_268 = x_214; +x_350 = lean_box(0); +if (lean_is_scalar(x_293)) { + x_351 = lean_alloc_ctor(0, 2, 0); } else { - lean_dec_ref(x_214); - x_268 = lean_box(0); + x_351 = x_293; } -x_269 = lean_box(0); -if (lean_is_scalar(x_268)) { - x_270 = lean_alloc_ctor(0, 2, 0); -} else { - x_270 = x_268; -} -lean_ctor_set(x_270, 0, x_269); -lean_ctor_set(x_270, 1, x_267); -return x_270; +lean_ctor_set(x_351, 0, x_350); +lean_ctor_set(x_351, 1, x_292); +return x_351; } } else { -lean_object* x_271; lean_object* x_272; -lean_dec(x_213); -x_271 = lean_ctor_get(x_214, 1); -lean_inc(x_271); -lean_dec(x_214); -x_272 = lean_box(0); -x_3 = x_272; -x_8 = x_203; -x_10 = x_271; +lean_object* x_352; lean_object* x_353; +lean_dec(x_288); +x_352 = lean_ctor_get(x_289, 1); +lean_inc(x_352); +lean_dec(x_289); +x_353 = lean_box(0); +x_3 = x_353; +x_8 = x_278; +x_10 = x_352; goto _start; } } else { -lean_object* x_274; lean_object* x_275; lean_object* x_276; lean_object* x_277; -lean_dec(x_213); -lean_dec(x_203); +lean_object* x_355; lean_object* x_356; lean_object* x_357; lean_object* x_358; +lean_dec(x_288); +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_274 = lean_ctor_get(x_214, 0); -lean_inc(x_274); -x_275 = lean_ctor_get(x_214, 1); -lean_inc(x_275); -if (lean_is_exclusive(x_214)) { - lean_ctor_release(x_214, 0); - lean_ctor_release(x_214, 1); - x_276 = x_214; +x_355 = lean_ctor_get(x_289, 0); +lean_inc(x_355); +x_356 = lean_ctor_get(x_289, 1); +lean_inc(x_356); +if (lean_is_exclusive(x_289)) { + lean_ctor_release(x_289, 0); + lean_ctor_release(x_289, 1); + x_357 = x_289; } else { - lean_dec_ref(x_214); - x_276 = lean_box(0); + lean_dec_ref(x_289); + x_357 = lean_box(0); } -if (lean_is_scalar(x_276)) { - x_277 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_357)) { + x_358 = lean_alloc_ctor(1, 2, 0); } else { - x_277 = x_276; + x_358 = x_357; } -lean_ctor_set(x_277, 0, x_274); -lean_ctor_set(x_277, 1, x_275); -return x_277; +lean_ctor_set(x_358, 0, x_355); +lean_ctor_set(x_358, 1, x_356); +return x_358; } } else { -lean_object* x_278; lean_object* x_279; -lean_dec(x_203); +lean_object* x_359; lean_object* x_360; +lean_dec(x_278); lean_dec(x_9); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_278 = lean_box(0); -if (lean_is_scalar(x_207)) { - x_279 = lean_alloc_ctor(0, 2, 0); +x_359 = lean_box(0); +if (lean_is_scalar(x_282)) { + x_360 = lean_alloc_ctor(0, 2, 0); } else { - x_279 = x_207; + x_360 = x_282; } -lean_ctor_set(x_279, 0, x_278); -lean_ctor_set(x_279, 1, x_206); -return x_279; +lean_ctor_set(x_360, 0, x_359); +lean_ctor_set(x_360, 1, x_281); +return x_360; } } else { -lean_object* x_280; -lean_dec(x_195); -lean_dec(x_194); -lean_dec(x_193); -lean_dec(x_192); -lean_dec(x_191); -lean_dec(x_189); -lean_dec(x_188); -lean_dec(x_187); -lean_dec(x_186); -lean_dec(x_185); -x_280 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(x_198, x_4, x_5, x_6, x_7, x_199, x_9, x_13); +lean_object* x_361; +lean_dec(x_270); +lean_dec(x_269); +lean_dec(x_268); +lean_dec(x_267); +lean_dec(x_266); +lean_dec(x_264); +lean_dec(x_263); +lean_dec(x_262); +lean_dec(x_261); +lean_dec(x_260); +x_361 = l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1(x_273, x_4, x_5, x_6, x_7, x_274, x_9, x_13); lean_dec(x_9); -lean_dec(x_199); +lean_dec(x_274); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_280; +return x_361; } } } else { -uint8_t x_281; +uint8_t x_362; 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_281 = !lean_is_exclusive(x_11); -if (x_281 == 0) +x_362 = !lean_is_exclusive(x_11); +if (x_362 == 0) { return x_11; } else { -lean_object* x_282; lean_object* x_283; lean_object* x_284; -x_282 = lean_ctor_get(x_11, 0); -x_283 = lean_ctor_get(x_11, 1); -lean_inc(x_283); -lean_inc(x_282); +lean_object* x_363; lean_object* x_364; lean_object* x_365; +x_363 = lean_ctor_get(x_11, 0); +x_364 = lean_ctor_get(x_11, 1); +lean_inc(x_364); +lean_inc(x_363); lean_dec(x_11); -x_284 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_284, 0, x_282); -lean_ctor_set(x_284, 1, x_283); -return x_284; +x_365 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_365, 0, x_363); +lean_ctor_set(x_365, 1, x_364); +return x_365; } } } @@ -17498,40 +18457,41 @@ lean_inc(x_5); x_16 = l_Lean_Elab_withInfoTreeContext___at_Lean_Elab_Term_runTactic___spec__9(x_2, x_15, 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_17; uint8_t x_18; lean_object* x_19; +lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); lean_dec(x_16); -x_18 = 0; -x_19 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_18, x_18, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -return x_19; +x_18 = 1; +x_19 = 0; +x_20 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_18, x_19, x_5, x_6, x_7, x_8, x_9, x_10, x_17); +return x_20; } else { -uint8_t x_20; +uint8_t x_21; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_20 = !lean_is_exclusive(x_16); -if (x_20 == 0) +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) { return x_16; } 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_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); lean_dec(x_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; +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; } } } @@ -18717,72 +19677,91 @@ lean_inc(x_3); x_11 = l_Lean_Elab_Term_synthesizeSyntheticMVars_loop(x_1, x_2, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_11) == 0) { -if (x_1 == 0) +uint8_t x_12; +x_12 = !lean_is_exclusive(x_11); +if (x_12 == 0) { -lean_object* x_12; lean_object* x_13; -x_12 = lean_ctor_get(x_11, 1); -lean_inc(x_12); -lean_dec(x_11); -x_13 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_processPostponedUniverseContraints(x_3, x_4, x_5, x_6, x_7, x_8, x_12); -return x_13; -} -else +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; +x_13 = lean_ctor_get(x_11, 1); +x_14 = lean_ctor_get(x_11, 0); +lean_dec(x_14); +x_15 = 1; +x_16 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_15); +if (x_16 == 0) { -uint8_t x_14; 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_14 = !lean_is_exclusive(x_11); -if (x_14 == 0) -{ -lean_object* x_15; -x_15 = lean_ctor_get(x_11, 0); -lean_dec(x_15); lean_ctor_set(x_11, 0, x_10); return x_11; } else { -lean_object* x_16; lean_object* x_17; -x_16 = lean_ctor_get(x_11, 1); -lean_inc(x_16); -lean_dec(x_11); -x_17 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_17, 0, x_10); -lean_ctor_set(x_17, 1, x_16); +lean_object* x_17; +lean_free_object(x_11); +x_17 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_processPostponedUniverseContraints(x_3, x_4, x_5, x_6, x_7, x_8, x_13); return x_17; } } +else +{ +lean_object* x_18; uint8_t x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_11, 1); +lean_inc(x_18); +lean_dec(x_11); +x_19 = 1; +x_20 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_1, x_19); +if (x_20 == 0) +{ +lean_object* x_21; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_10); +lean_ctor_set(x_21, 1, x_18); +return x_21; +} +else +{ +lean_object* x_22; +x_22 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_processPostponedUniverseContraints(x_3, x_4, x_5, x_6, x_7, x_8, x_18); +return x_22; +} +} } else { -uint8_t x_18; +uint8_t x_23; 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_18 = !lean_is_exclusive(x_11); -if (x_18 == 0) +x_23 = !lean_is_exclusive(x_11); +if (x_23 == 0) { return x_11; } else { -lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_19 = lean_ctor_get(x_11, 0); -x_20 = lean_ctor_get(x_11, 1); -lean_inc(x_20); -lean_inc(x_19); +lean_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_inc(x_25); +lean_inc(x_24); lean_dec(x_11); -x_21 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_21, 0, x_19); -lean_ctor_set(x_21, 1, x_20); -return x_21; +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; } } } @@ -19009,7 +19988,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsNoPostponing(u _start: { uint8_t x_9; lean_object* x_10; -x_9 = 0; +x_9 = 1; x_10 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_9, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); return x_10; } @@ -19079,7 +20058,7 @@ lean_object* x_17; uint8_t x_18; uint8_t x_19; lean_object* x_20; x_17 = lean_ctor_get(x_8, 1); lean_inc(x_17); lean_dec(x_8); -x_18 = 1; +x_18 = 0; x_19 = 0; lean_inc(x_6); lean_inc(x_5); @@ -19161,7 +20140,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeSyntheticMVarsUsingDefault(l _start: { uint8_t x_8; uint8_t x_9; lean_object* x_10; -x_8 = 1; +x_8 = 0; x_9 = 0; lean_inc(x_6); lean_inc(x_5); @@ -19181,339 +20160,827 @@ return x_12; } else { -uint8_t x_13; -lean_dec(x_6); -lean_dec(x_5); +uint8_t x_13; +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +lean_dec(x_1); +x_13 = !lean_is_exclusive(x_10); +if (x_13 == 0) +{ +return x_10; +} +else +{ +lean_object* x_14; lean_object* x_15; lean_object* x_16; +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); +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_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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) { +_start: +{ +lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_10 = lean_box(0); +x_11 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_11, 0, x_1); +lean_ctor_set(x_11, 1, x_10); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_9); +return x_12; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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, 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; uint8_t x_17; +x_10 = lean_st_ref_get(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_ctor_get(x_11, 2); +lean_inc(x_13); +lean_dec(x_11); +x_14 = lean_st_ref_take(x_4, 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_is_exclusive(x_15); +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_50; lean_object* x_51; lean_object* x_75; +x_18 = lean_ctor_get(x_15, 2); +lean_dec(x_18); +x_19 = lean_box(0); +lean_ctor_set(x_15, 2, x_19); +x_20 = lean_st_ref_set(x_4, x_15, x_16); +x_21 = lean_ctor_get(x_20, 1); +lean_inc(x_21); +lean_dec(x_20); +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_75 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_21); +if (lean_obj_tag(x_75) == 0) +{ +lean_object* x_76; lean_object* x_77; uint8_t x_78; lean_object* x_79; +x_76 = lean_ctor_get(x_75, 0); +lean_inc(x_76); +x_77 = lean_ctor_get(x_75, 1); +lean_inc(x_77); +lean_dec(x_75); +x_78 = 0; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_79 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_2, x_78, x_3, x_4, x_5, x_6, x_7, x_8, x_77); +if (lean_obj_tag(x_79) == 0) +{ +lean_object* x_80; uint8_t x_81; uint8_t x_82; +x_80 = lean_ctor_get(x_79, 1); +lean_inc(x_80); +lean_dec(x_79); +x_81 = 0; +x_82 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_2, x_81); +if (x_82 == 0) +{ +lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; +x_83 = lean_box(0); +x_84 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_76, x_83, x_3, x_4, x_5, x_6, x_7, x_8, x_80); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +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_22 = x_85; +x_23 = x_86; +goto block_49; +} +else +{ +lean_object* x_87; +lean_inc(x_8); +lean_inc(x_7); +lean_inc(x_6); +lean_inc(x_5); +lean_inc(x_4); +lean_inc(x_3); +x_87 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultLoop(x_3, x_4, x_5, x_6, x_7, x_8, x_80); +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; +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___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_76, x_88, x_3, x_4, x_5, x_6, x_7, x_8, x_89); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_88); +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_22 = x_91; +x_23 = x_92; +goto block_49; +} +else +{ +lean_object* x_93; lean_object* x_94; +lean_dec(x_76); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_93 = lean_ctor_get(x_87, 0); +lean_inc(x_93); +x_94 = lean_ctor_get(x_87, 1); +lean_inc(x_94); +lean_dec(x_87); +x_50 = x_93; +x_51 = x_94; +goto block_74; +} +} +} +else +{ +lean_object* x_95; lean_object* x_96; +lean_dec(x_76); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_95 = lean_ctor_get(x_79, 0); +lean_inc(x_95); +x_96 = lean_ctor_get(x_79, 1); +lean_inc(x_96); +lean_dec(x_79); +x_50 = x_95; +x_51 = x_96; +goto block_74; +} +} +else +{ +lean_object* x_97; lean_object* x_98; +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_97 = lean_ctor_get(x_75, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_75, 1); +lean_inc(x_98); +lean_dec(x_75); +x_50 = x_97; +x_51 = x_98; +goto block_74; +} +block_49: +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; uint8_t x_27; +x_24 = lean_st_ref_take(x_4, x_23); +x_25 = lean_ctor_get(x_24, 0); +lean_inc(x_25); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = !lean_is_exclusive(x_25); +if (x_27 == 0) +{ +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; +x_28 = lean_ctor_get(x_25, 2); +x_29 = l_List_appendTR___rarg(x_28, x_13); +lean_ctor_set(x_25, 2, x_29); +x_30 = lean_st_ref_set(x_4, x_25, x_26); +lean_dec(x_4); +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, 0); +lean_dec(x_32); +x_33 = lean_ctor_get(x_22, 0); +lean_inc(x_33); +lean_dec(x_22); +lean_ctor_set(x_30, 0, x_33); +return x_30; +} +else +{ +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_30, 1); +lean_inc(x_34); +lean_dec(x_30); +x_35 = lean_ctor_get(x_22, 0); +lean_inc(x_35); +lean_dec(x_22); +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; +} +} +else +{ +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; +x_37 = lean_ctor_get(x_25, 0); +x_38 = lean_ctor_get(x_25, 1); +x_39 = lean_ctor_get(x_25, 2); +x_40 = lean_ctor_get(x_25, 3); +x_41 = lean_ctor_get(x_25, 4); +lean_inc(x_41); +lean_inc(x_40); +lean_inc(x_39); +lean_inc(x_38); +lean_inc(x_37); +lean_dec(x_25); +x_42 = l_List_appendTR___rarg(x_39, x_13); +x_43 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_43, 0, x_37); +lean_ctor_set(x_43, 1, x_38); +lean_ctor_set(x_43, 2, x_42); +lean_ctor_set(x_43, 3, x_40); +lean_ctor_set(x_43, 4, x_41); +x_44 = lean_st_ref_set(x_4, x_43, x_26); +lean_dec(x_4); +x_45 = lean_ctor_get(x_44, 1); +lean_inc(x_45); +if (lean_is_exclusive(x_44)) { + lean_ctor_release(x_44, 0); + lean_ctor_release(x_44, 1); + x_46 = x_44; +} else { + lean_dec_ref(x_44); + x_46 = lean_box(0); +} +x_47 = lean_ctor_get(x_22, 0); +lean_inc(x_47); +lean_dec(x_22); +if (lean_is_scalar(x_46)) { + x_48 = lean_alloc_ctor(0, 2, 0); +} else { + x_48 = x_46; +} +lean_ctor_set(x_48, 0, x_47); +lean_ctor_set(x_48, 1, x_45); +return x_48; +} +} +block_74: +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_52 = lean_st_ref_take(x_4, x_51); +x_53 = lean_ctor_get(x_52, 0); +lean_inc(x_53); +x_54 = lean_ctor_get(x_52, 1); +lean_inc(x_54); +lean_dec(x_52); +x_55 = !lean_is_exclusive(x_53); +if (x_55 == 0) +{ +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; +x_56 = lean_ctor_get(x_53, 2); +x_57 = l_List_appendTR___rarg(x_56, x_13); +lean_ctor_set(x_53, 2, x_57); +x_58 = lean_st_ref_set(x_4, x_53, x_54); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -x_13 = !lean_is_exclusive(x_10); -if (x_13 == 0) +x_59 = !lean_is_exclusive(x_58); +if (x_59 == 0) { -return x_10; +lean_object* x_60; +x_60 = lean_ctor_get(x_58, 0); +lean_dec(x_60); +lean_ctor_set_tag(x_58, 1); +lean_ctor_set(x_58, 0, x_50); +return x_58; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -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); -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_object* x_61; lean_object* x_62; +x_61 = lean_ctor_get(x_58, 1); +lean_inc(x_61); +lean_dec(x_58); +x_62 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_62, 0, x_50); +lean_ctor_set(x_62, 1, x_61); +return x_62; +} } +else +{ +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_object* x_72; lean_object* x_73; +x_63 = lean_ctor_get(x_53, 0); +x_64 = lean_ctor_get(x_53, 1); +x_65 = lean_ctor_get(x_53, 2); +x_66 = lean_ctor_get(x_53, 3); +x_67 = lean_ctor_get(x_53, 4); +lean_inc(x_67); +lean_inc(x_66); +lean_inc(x_65); +lean_inc(x_64); +lean_inc(x_63); +lean_dec(x_53); +x_68 = l_List_appendTR___rarg(x_65, x_13); +x_69 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_69, 0, x_63); +lean_ctor_set(x_69, 1, x_64); +lean_ctor_set(x_69, 2, x_68); +lean_ctor_set(x_69, 3, x_66); +lean_ctor_set(x_69, 4, x_67); +x_70 = lean_st_ref_set(x_4, x_69, x_54); +lean_dec(x_4); +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); } +if (lean_is_scalar(x_72)) { + x_73 = lean_alloc_ctor(1, 2, 0); +} else { + x_73 = x_72; + lean_ctor_set_tag(x_73, 1); } +lean_ctor_set(x_73, 0, x_50); +lean_ctor_set(x_73, 1, x_71); +return x_73; } -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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) { -_start: -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_10 = lean_box(0); -x_11 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_11, 0, x_1); -lean_ctor_set(x_11, 1, x_10); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_9); -return x_12; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(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, lean_object* x_9, lean_object* x_10) { -_start: +else { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_11 = lean_st_ref_get(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_ctor_get(x_12, 2); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_st_ref_take(x_5, x_13); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); +lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_126; lean_object* x_127; lean_object* x_144; +x_99 = lean_ctor_get(x_15, 0); +x_100 = lean_ctor_get(x_15, 1); +x_101 = lean_ctor_get(x_15, 3); +x_102 = lean_ctor_get(x_15, 4); +lean_inc(x_102); +lean_inc(x_101); +lean_inc(x_100); +lean_inc(x_99); 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; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_51; lean_object* x_52; lean_object* x_76; -x_19 = lean_ctor_get(x_16, 2); -lean_dec(x_19); -x_20 = lean_box(0); -lean_ctor_set(x_16, 2, x_20); -x_21 = lean_st_ref_set(x_5, x_16, x_17); -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); -lean_inc(x_9); +x_103 = lean_box(0); +x_104 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_104, 0, x_99); +lean_ctor_set(x_104, 1, x_100); +lean_ctor_set(x_104, 2, x_103); +lean_ctor_set(x_104, 3, x_101); +lean_ctor_set(x_104, 4, x_102); +x_105 = lean_st_ref_set(x_4, x_104, x_16); +x_106 = lean_ctor_get(x_105, 1); +lean_inc(x_106); +lean_dec(x_105); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_76 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_22); -if (lean_obj_tag(x_76) == 0) +lean_inc(x_3); +x_144 = lean_apply_7(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_106); +if (lean_obj_tag(x_144) == 0) { -lean_object* x_77; lean_object* x_78; uint8_t x_79; lean_object* x_80; -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 = 0; -lean_inc(x_9); +lean_object* x_145; lean_object* x_146; uint8_t x_147; lean_object* x_148; +x_145 = lean_ctor_get(x_144, 0); +lean_inc(x_145); +x_146 = lean_ctor_get(x_144, 1); +lean_inc(x_146); +lean_dec(x_144); +x_147 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_80 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_2, x_79, x_4, x_5, x_6, x_7, x_8, x_9, x_78); -if (lean_obj_tag(x_80) == 0) -{ -if (x_2 == 0) -{ -lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; -x_81 = lean_ctor_get(x_80, 1); -lean_inc(x_81); -lean_dec(x_80); -x_82 = lean_box(0); -x_83 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_77, x_82, x_4, x_5, x_6, x_7, x_8, x_9, x_81); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -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_23 = x_84; -x_24 = x_85; -goto block_50; -} -else +lean_inc(x_3); +x_148 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_2, x_147, x_3, x_4, x_5, x_6, x_7, x_8, x_146); +if (lean_obj_tag(x_148) == 0) { -if (x_3 == 0) +lean_object* x_149; uint8_t x_150; uint8_t x_151; +x_149 = lean_ctor_get(x_148, 1); +lean_inc(x_149); +lean_dec(x_148); +x_150 = 0; +x_151 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_beqPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_4101_(x_2, x_150); +if (x_151 == 0) { -lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; -x_86 = lean_ctor_get(x_80, 1); -lean_inc(x_86); -lean_dec(x_80); -x_87 = lean_box(0); -x_88 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_77, x_87, x_4, x_5, x_6, x_7, x_8, x_9, x_86); -lean_dec(x_9); +lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; +x_152 = lean_box(0); +x_153 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_145, x_152, x_3, x_4, x_5, x_6, x_7, x_8, x_149); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -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_23 = x_89; -x_24 = x_90; -goto block_50; +lean_dec(x_5); +lean_dec(x_3); +x_154 = lean_ctor_get(x_153, 0); +lean_inc(x_154); +x_155 = lean_ctor_get(x_153, 1); +lean_inc(x_155); +lean_dec(x_153); +x_107 = x_154; +x_108 = x_155; +goto block_125; } else { -lean_object* x_91; lean_object* x_92; -x_91 = lean_ctor_get(x_80, 1); -lean_inc(x_91); -lean_dec(x_80); -lean_inc(x_9); +lean_object* x_156; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_92 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultLoop(x_4, x_5, x_6, x_7, x_8, x_9, x_91); -if (lean_obj_tag(x_92) == 0) +lean_inc(x_3); +x_156 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultLoop(x_3, x_4, x_5, x_6, x_7, x_8, x_149); +if (lean_obj_tag(x_156) == 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 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_77, x_93, x_4, x_5, x_6, x_7, x_8, x_9, x_94); -lean_dec(x_9); +lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; lean_object* x_161; +x_157 = lean_ctor_get(x_156, 0); +lean_inc(x_157); +x_158 = lean_ctor_get(x_156, 1); +lean_inc(x_158); +lean_dec(x_156); +x_159 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_145, x_157, x_3, x_4, x_5, x_6, x_7, x_8, x_158); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_93); -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_23 = x_96; -x_24 = x_97; -goto block_50; +lean_dec(x_5); +lean_dec(x_3); +lean_dec(x_157); +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_107 = x_160; +x_108 = x_161; +goto block_125; } else { -lean_object* x_98; lean_object* x_99; -lean_dec(x_77); -lean_dec(x_9); +lean_object* x_162; lean_object* x_163; +lean_dec(x_145); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); -lean_dec(x_4); -x_98 = lean_ctor_get(x_92, 0); -lean_inc(x_98); -x_99 = lean_ctor_get(x_92, 1); -lean_inc(x_99); -lean_dec(x_92); -x_51 = x_98; -x_52 = x_99; -goto block_75; +lean_dec(x_5); +lean_dec(x_3); +x_162 = lean_ctor_get(x_156, 0); +lean_inc(x_162); +x_163 = lean_ctor_get(x_156, 1); +lean_inc(x_163); +lean_dec(x_156); +x_126 = x_162; +x_127 = x_163; +goto block_143; } } } +else +{ +lean_object* x_164; lean_object* x_165; +lean_dec(x_145); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_164 = lean_ctor_get(x_148, 0); +lean_inc(x_164); +x_165 = lean_ctor_get(x_148, 1); +lean_inc(x_165); +lean_dec(x_148); +x_126 = x_164; +x_127 = x_165; +goto block_143; +} } else { -lean_object* x_100; lean_object* x_101; -lean_dec(x_77); -lean_dec(x_9); +lean_object* x_166; lean_object* x_167; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_3); +x_166 = lean_ctor_get(x_144, 0); +lean_inc(x_166); +x_167 = lean_ctor_get(x_144, 1); +lean_inc(x_167); +lean_dec(x_144); +x_126 = x_166; +x_127 = x_167; +goto block_143; +} +block_125: +{ +lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; +x_109 = lean_st_ref_take(x_4, x_108); +x_110 = lean_ctor_get(x_109, 0); +lean_inc(x_110); +x_111 = lean_ctor_get(x_109, 1); +lean_inc(x_111); +lean_dec(x_109); +x_112 = lean_ctor_get(x_110, 0); +lean_inc(x_112); +x_113 = lean_ctor_get(x_110, 1); +lean_inc(x_113); +x_114 = lean_ctor_get(x_110, 2); +lean_inc(x_114); +x_115 = lean_ctor_get(x_110, 3); +lean_inc(x_115); +x_116 = lean_ctor_get(x_110, 4); +lean_inc(x_116); +if (lean_is_exclusive(x_110)) { + lean_ctor_release(x_110, 0); + lean_ctor_release(x_110, 1); + lean_ctor_release(x_110, 2); + lean_ctor_release(x_110, 3); + lean_ctor_release(x_110, 4); + x_117 = x_110; +} else { + lean_dec_ref(x_110); + x_117 = lean_box(0); +} +x_118 = l_List_appendTR___rarg(x_114, x_13); +if (lean_is_scalar(x_117)) { + x_119 = lean_alloc_ctor(0, 5, 0); +} else { + x_119 = x_117; +} +lean_ctor_set(x_119, 0, x_112); +lean_ctor_set(x_119, 1, x_113); +lean_ctor_set(x_119, 2, x_118); +lean_ctor_set(x_119, 3, x_115); +lean_ctor_set(x_119, 4, x_116); +x_120 = lean_st_ref_set(x_4, x_119, x_111); lean_dec(x_4); -x_100 = lean_ctor_get(x_80, 0); -lean_inc(x_100); -x_101 = lean_ctor_get(x_80, 1); -lean_inc(x_101); -lean_dec(x_80); -x_51 = x_100; -x_52 = x_101; -goto block_75; +x_121 = lean_ctor_get(x_120, 1); +lean_inc(x_121); +if (lean_is_exclusive(x_120)) { + lean_ctor_release(x_120, 0); + lean_ctor_release(x_120, 1); + x_122 = x_120; +} else { + lean_dec_ref(x_120); + x_122 = lean_box(0); +} +x_123 = lean_ctor_get(x_107, 0); +lean_inc(x_123); +lean_dec(x_107); +if (lean_is_scalar(x_122)) { + x_124 = lean_alloc_ctor(0, 2, 0); +} else { + x_124 = x_122; +} +lean_ctor_set(x_124, 0, x_123); +lean_ctor_set(x_124, 1, x_121); +return x_124; +} +block_143: +{ +lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; +x_128 = lean_st_ref_take(x_4, x_127); +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_inc(x_131); +x_132 = lean_ctor_get(x_129, 1); +lean_inc(x_132); +x_133 = lean_ctor_get(x_129, 2); +lean_inc(x_133); +x_134 = lean_ctor_get(x_129, 3); +lean_inc(x_134); +x_135 = lean_ctor_get(x_129, 4); +lean_inc(x_135); +if (lean_is_exclusive(x_129)) { + lean_ctor_release(x_129, 0); + lean_ctor_release(x_129, 1); + lean_ctor_release(x_129, 2); + lean_ctor_release(x_129, 3); + lean_ctor_release(x_129, 4); + x_136 = x_129; +} else { + lean_dec_ref(x_129); + x_136 = lean_box(0); +} +x_137 = l_List_appendTR___rarg(x_133, x_13); +if (lean_is_scalar(x_136)) { + x_138 = lean_alloc_ctor(0, 5, 0); +} else { + x_138 = x_136; +} +lean_ctor_set(x_138, 0, x_131); +lean_ctor_set(x_138, 1, x_132); +lean_ctor_set(x_138, 2, x_137); +lean_ctor_set(x_138, 3, x_134); +lean_ctor_set(x_138, 4, x_135); +x_139 = lean_st_ref_set(x_4, x_138, x_130); +lean_dec(x_4); +x_140 = lean_ctor_get(x_139, 1); +lean_inc(x_140); +if (lean_is_exclusive(x_139)) { + lean_ctor_release(x_139, 0); + lean_ctor_release(x_139, 1); + x_141 = x_139; +} else { + lean_dec_ref(x_139); + x_141 = lean_box(0); +} +if (lean_is_scalar(x_141)) { + x_142 = lean_alloc_ctor(1, 2, 0); +} else { + x_142 = x_141; + lean_ctor_set_tag(x_142, 1); +} +lean_ctor_set(x_142, 0, x_126); +lean_ctor_set(x_142, 1, x_140); +return x_142; +} +} } } -else +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp(lean_object* x_1) { +_start: { -lean_object* x_102; lean_object* x_103; -lean_dec(x_9); +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 0); +return x_2; +} +} +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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) { +_start: +{ +lean_object* x_10; +x_10 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -x_102 = lean_ctor_get(x_76, 0); -lean_inc(x_102); -x_103 = lean_ctor_get(x_76, 1); -lean_inc(x_103); -lean_dec(x_76); -x_51 = x_102; -x_52 = x_103; -goto block_75; +lean_dec(x_3); +lean_dec(x_2); +return x_10; +} } -block_50: +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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_25; lean_object* x_26; lean_object* x_27; uint8_t x_28; -x_25 = lean_st_ref_take(x_5, 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) +uint8_t x_10; lean_object* x_11; +x_10 = lean_unbox(x_2); +lean_dec(x_2); +x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___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, lean_object* x_9, lean_object* x_10) { +_start: { -lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_29 = lean_ctor_get(x_26, 2); -x_30 = l_List_appendTR___rarg(x_29, x_14); -lean_ctor_set(x_26, 2, x_30); -x_31 = lean_st_ref_set(x_5, x_26, x_27); -lean_dec(x_5); -x_32 = !lean_is_exclusive(x_31); -if (x_32 == 0) +lean_object* x_11; +x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_3, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_11; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { +_start: { -lean_object* x_33; lean_object* x_34; -x_33 = lean_ctor_get(x_31, 0); -lean_dec(x_33); -x_34 = lean_ctor_get(x_23, 0); -lean_inc(x_34); -lean_dec(x_23); -lean_ctor_set(x_31, 0, x_34); -return x_31; +lean_object* x_5; lean_object* x_6; lean_object* x_7; +x_5 = lean_box(x_4); +x_6 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withSynthesize___rarg___lambda__1___boxed), 10, 1); +lean_closure_set(x_6, 0, x_5); +x_7 = lean_apply_3(x_1, lean_box(0), x_6, x_3); +return x_7; } -else +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize(lean_object* x_1, lean_object* x_2) { +_start: { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_31, 1); -lean_inc(x_35); -lean_dec(x_31); -x_36 = lean_ctor_get(x_23, 0); -lean_inc(x_36); -lean_dec(x_23); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_35); -return x_37; +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withSynthesize___rarg___boxed), 4, 0); +return x_3; } } -else +LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___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) { +_start: { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; 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_26, 0); -x_39 = lean_ctor_get(x_26, 1); -x_40 = lean_ctor_get(x_26, 2); -x_41 = lean_ctor_get(x_26, 3); -x_42 = lean_ctor_get(x_26, 4); -lean_inc(x_42); -lean_inc(x_41); -lean_inc(x_40); -lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_26); -x_43 = l_List_appendTR___rarg(x_40, x_14); -x_44 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_44, 0, x_38); -lean_ctor_set(x_44, 1, x_39); -lean_ctor_set(x_44, 2, x_43); -lean_ctor_set(x_44, 3, x_41); -lean_ctor_set(x_44, 4, x_42); -x_45 = lean_st_ref_set(x_5, x_44, x_27); -lean_dec(x_5); -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); +uint8_t x_11; lean_object* x_12; +x_11 = lean_unbox(x_1); +lean_dec(x_1); +x_12 = l_Lean_Elab_Term_withSynthesize___rarg___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_12; } -x_48 = lean_ctor_get(x_23, 0); -lean_inc(x_48); -lean_dec(x_23); -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_Lean_Elab_Term_withSynthesize___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +_start: +{ +uint8_t x_5; lean_object* x_6; +x_5 = lean_unbox(x_4); +lean_dec(x_4); +x_6 = l_Lean_Elab_Term_withSynthesize___rarg(x_1, x_2, x_3, x_5); +lean_dec(x_2); +return x_6; } } -block_75: +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_9 = lean_st_ref_get(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_ctor_get(x_10, 2); +lean_inc(x_12); +lean_dec(x_10); +x_13 = lean_st_ref_take(x_3, x_11); +x_14 = lean_ctor_get(x_13, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_13, 1); +lean_inc(x_15); +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; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_46; +x_17 = lean_ctor_get(x_14, 2); +lean_dec(x_17); +x_18 = lean_box(0); +lean_ctor_set(x_14, 2, x_18); +x_19 = lean_st_ref_set(x_3, x_14, x_15); +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +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_46 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_20); +if (lean_obj_tag(x_46) == 0) +{ +lean_object* x_47; lean_object* x_48; uint8_t x_49; uint8_t 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 = 0; +x_50 = 0; +lean_inc(x_3); +x_51 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_49, x_50, x_2, x_3, x_4, x_5, x_6, x_7, x_48); +if (lean_obj_tag(x_51) == 0) { -lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; -x_53 = lean_st_ref_take(x_5, x_52); +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +x_52 = lean_ctor_get(x_51, 1); +lean_inc(x_52); +lean_dec(x_51); +x_53 = lean_st_ref_take(x_3, x_52); x_54 = lean_ctor_get(x_53, 0); lean_inc(x_54); x_55 = lean_ctor_get(x_53, 1); @@ -19524,18 +20991,17 @@ if (x_56 == 0) { lean_object* x_57; lean_object* x_58; lean_object* x_59; uint8_t x_60; x_57 = lean_ctor_get(x_54, 2); -x_58 = l_List_appendTR___rarg(x_57, x_14); +x_58 = l_List_appendTR___rarg(x_57, x_12); lean_ctor_set(x_54, 2, x_58); -x_59 = lean_st_ref_set(x_5, x_54, x_55); -lean_dec(x_5); +x_59 = lean_st_ref_set(x_3, x_54, x_55); +lean_dec(x_3); x_60 = !lean_is_exclusive(x_59); if (x_60 == 0) { lean_object* x_61; x_61 = lean_ctor_get(x_59, 0); lean_dec(x_61); -lean_ctor_set_tag(x_59, 1); -lean_ctor_set(x_59, 0, x_51); +lean_ctor_set(x_59, 0, x_47); return x_59; } else @@ -19544,8 +21010,8 @@ lean_object* x_62; lean_object* x_63; x_62 = lean_ctor_get(x_59, 1); lean_inc(x_62); lean_dec(x_59); -x_63 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_63, 0, x_51); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_47); lean_ctor_set(x_63, 1, x_62); return x_63; } @@ -19564,15 +21030,15 @@ lean_inc(x_66); lean_inc(x_65); lean_inc(x_64); lean_dec(x_54); -x_69 = l_List_appendTR___rarg(x_66, x_14); +x_69 = l_List_appendTR___rarg(x_66, x_12); x_70 = lean_alloc_ctor(0, 5, 0); lean_ctor_set(x_70, 0, x_64); lean_ctor_set(x_70, 1, x_65); lean_ctor_set(x_70, 2, x_69); lean_ctor_set(x_70, 3, x_67); lean_ctor_set(x_70, 4, x_68); -x_71 = lean_st_ref_set(x_5, x_70, x_55); -lean_dec(x_5); +x_71 = lean_st_ref_set(x_3, x_70, x_55); +lean_dec(x_3); x_72 = lean_ctor_get(x_71, 1); lean_inc(x_72); if (lean_is_exclusive(x_71)) { @@ -19584,428 +21050,349 @@ if (lean_is_exclusive(x_71)) { x_73 = lean_box(0); } if (lean_is_scalar(x_73)) { - x_74 = lean_alloc_ctor(1, 2, 0); + x_74 = lean_alloc_ctor(0, 2, 0); } else { x_74 = x_73; - lean_ctor_set_tag(x_74, 1); } -lean_ctor_set(x_74, 0, x_51); +lean_ctor_set(x_74, 0, x_47); lean_ctor_set(x_74, 1, x_72); return x_74; } } -} else { -lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_131; lean_object* x_132; lean_object* x_149; -x_104 = lean_ctor_get(x_16, 0); -x_105 = lean_ctor_get(x_16, 1); -x_106 = lean_ctor_get(x_16, 3); -x_107 = lean_ctor_get(x_16, 4); -lean_inc(x_107); -lean_inc(x_106); -lean_inc(x_105); -lean_inc(x_104); -lean_dec(x_16); -x_108 = lean_box(0); -x_109 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_109, 0, x_104); -lean_ctor_set(x_109, 1, x_105); -lean_ctor_set(x_109, 2, x_108); -lean_ctor_set(x_109, 3, x_106); -lean_ctor_set(x_109, 4, x_107); -x_110 = lean_st_ref_set(x_5, x_109, x_17); -x_111 = lean_ctor_get(x_110, 1); -lean_inc(x_111); -lean_dec(x_110); -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_149 = lean_apply_7(x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_111); -if (lean_obj_tag(x_149) == 0) -{ -lean_object* x_150; lean_object* x_151; uint8_t x_152; lean_object* x_153; -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 = 0; -lean_inc(x_9); -lean_inc(x_8); -lean_inc(x_7); -lean_inc(x_6); -lean_inc(x_5); -lean_inc(x_4); -x_153 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_2, x_152, x_4, x_5, x_6, x_7, x_8, x_9, x_151); -if (lean_obj_tag(x_153) == 0) -{ -if (x_2 == 0) -{ -lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; -x_154 = lean_ctor_get(x_153, 1); -lean_inc(x_154); -lean_dec(x_153); -x_155 = lean_box(0); -x_156 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_150, x_155, x_4, x_5, x_6, x_7, x_8, x_9, x_154); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_157 = lean_ctor_get(x_156, 0); -lean_inc(x_157); -x_158 = lean_ctor_get(x_156, 1); -lean_inc(x_158); -lean_dec(x_156); -x_112 = x_157; -x_113 = x_158; -goto block_130; +lean_object* x_75; lean_object* x_76; +lean_dec(x_47); +x_75 = lean_ctor_get(x_51, 0); +lean_inc(x_75); +x_76 = lean_ctor_get(x_51, 1); +lean_inc(x_76); +lean_dec(x_51); +x_21 = x_75; +x_22 = x_76; +goto block_45; +} } else { -if (x_3 == 0) -{ -lean_object* x_159; lean_object* x_160; lean_object* x_161; lean_object* x_162; lean_object* x_163; -x_159 = lean_ctor_get(x_153, 1); -lean_inc(x_159); -lean_dec(x_153); -x_160 = lean_box(0); -x_161 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_150, x_160, x_4, x_5, x_6, x_7, x_8, x_9, x_159); -lean_dec(x_9); -lean_dec(x_8); +lean_object* x_77; lean_object* x_78; lean_dec(x_7); lean_dec(x_6); +lean_dec(x_5); lean_dec(x_4); -x_162 = lean_ctor_get(x_161, 0); -lean_inc(x_162); -x_163 = lean_ctor_get(x_161, 1); -lean_inc(x_163); -lean_dec(x_161); -x_112 = x_162; -x_113 = x_163; -goto block_130; +lean_dec(x_2); +x_77 = lean_ctor_get(x_46, 0); +lean_inc(x_77); +x_78 = lean_ctor_get(x_46, 1); +lean_inc(x_78); +lean_dec(x_46); +x_21 = x_77; +x_22 = x_78; +goto block_45; } -else +block_45: { -lean_object* x_164; lean_object* x_165; -x_164 = lean_ctor_get(x_153, 1); -lean_inc(x_164); -lean_dec(x_153); -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_165 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultLoop(x_4, x_5, x_6, x_7, x_8, x_9, x_164); -if (lean_obj_tag(x_165) == 0) +lean_object* x_23; lean_object* x_24; lean_object* x_25; uint8_t x_26; +x_23 = lean_st_ref_take(x_3, x_22); +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); +lean_dec(x_23); +x_26 = !lean_is_exclusive(x_24); +if (x_26 == 0) { -lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; -x_166 = lean_ctor_get(x_165, 0); -lean_inc(x_166); -x_167 = lean_ctor_get(x_165, 1); -lean_inc(x_167); -lean_dec(x_165); -x_168 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_150, x_166, x_4, x_5, x_6, x_7, x_8, x_9, x_167); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -lean_dec(x_166); -x_169 = lean_ctor_get(x_168, 0); -lean_inc(x_169); -x_170 = lean_ctor_get(x_168, 1); -lean_inc(x_170); -lean_dec(x_168); -x_112 = x_169; -x_113 = x_170; -goto block_130; +lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; +x_27 = lean_ctor_get(x_24, 2); +x_28 = l_List_appendTR___rarg(x_27, x_12); +lean_ctor_set(x_24, 2, x_28); +x_29 = lean_st_ref_set(x_3, x_24, x_25); +lean_dec(x_3); +x_30 = !lean_is_exclusive(x_29); +if (x_30 == 0) +{ +lean_object* x_31; +x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_31); +lean_ctor_set_tag(x_29, 1); +lean_ctor_set(x_29, 0, x_21); +return x_29; } else { -lean_object* x_171; lean_object* x_172; -lean_dec(x_150); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_171 = lean_ctor_get(x_165, 0); -lean_inc(x_171); -x_172 = lean_ctor_get(x_165, 1); -lean_inc(x_172); -lean_dec(x_165); -x_131 = x_171; -x_132 = x_172; -goto block_148; -} -} +lean_object* x_32; lean_object* x_33; +x_32 = lean_ctor_get(x_29, 1); +lean_inc(x_32); +lean_dec(x_29); +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_21); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } else { -lean_object* x_173; lean_object* x_174; -lean_dec(x_150); -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_173 = lean_ctor_get(x_153, 0); -lean_inc(x_173); -x_174 = lean_ctor_get(x_153, 1); -lean_inc(x_174); -lean_dec(x_153); -x_131 = x_173; -x_132 = x_174; -goto block_148; +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; 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_ctor_get(x_24, 0); +x_35 = lean_ctor_get(x_24, 1); +x_36 = lean_ctor_get(x_24, 2); +x_37 = lean_ctor_get(x_24, 3); +x_38 = lean_ctor_get(x_24, 4); +lean_inc(x_38); +lean_inc(x_37); +lean_inc(x_36); +lean_inc(x_35); +lean_inc(x_34); +lean_dec(x_24); +x_39 = l_List_appendTR___rarg(x_36, x_12); +x_40 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_40, 0, x_34); +lean_ctor_set(x_40, 1, x_35); +lean_ctor_set(x_40, 2, x_39); +lean_ctor_set(x_40, 3, x_37); +lean_ctor_set(x_40, 4, x_38); +x_41 = lean_st_ref_set(x_3, x_40, x_25); +lean_dec(x_3); +x_42 = lean_ctor_get(x_41, 1); +lean_inc(x_42); +if (lean_is_exclusive(x_41)) { + lean_ctor_release(x_41, 0); + lean_ctor_release(x_41, 1); + x_43 = x_41; +} else { + lean_dec_ref(x_41); + x_43 = lean_box(0); +} +if (lean_is_scalar(x_43)) { + x_44 = lean_alloc_ctor(1, 2, 0); +} else { + x_44 = x_43; + lean_ctor_set_tag(x_44, 1); +} +lean_ctor_set(x_44, 0, x_21); +lean_ctor_set(x_44, 1, x_42); +return x_44; +} } } else { -lean_object* x_175; lean_object* x_176; -lean_dec(x_9); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_4); -x_175 = lean_ctor_get(x_149, 0); -lean_inc(x_175); -x_176 = lean_ctor_get(x_149, 1); -lean_inc(x_176); -lean_dec(x_149); -x_131 = x_175; -x_132 = x_176; -goto block_148; -} -block_130: +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_105; +x_79 = lean_ctor_get(x_14, 0); +x_80 = lean_ctor_get(x_14, 1); +x_81 = lean_ctor_get(x_14, 3); +x_82 = lean_ctor_get(x_14, 4); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_inc(x_79); +lean_dec(x_14); +x_83 = lean_box(0); +x_84 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_84, 0, x_79); +lean_ctor_set(x_84, 1, x_80); +lean_ctor_set(x_84, 2, x_83); +lean_ctor_set(x_84, 3, x_81); +lean_ctor_set(x_84, 4, x_82); +x_85 = lean_st_ref_set(x_3, x_84, x_15); +x_86 = lean_ctor_get(x_85, 1); +lean_inc(x_86); +lean_dec(x_85); +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_105 = lean_apply_7(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_86); +if (lean_obj_tag(x_105) == 0) { -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; -x_114 = lean_st_ref_take(x_5, x_113); -x_115 = lean_ctor_get(x_114, 0); +lean_object* x_106; lean_object* x_107; uint8_t x_108; uint8_t x_109; lean_object* x_110; +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 = 0; +x_109 = 0; +lean_inc(x_3); +x_110 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_108, x_109, x_2, x_3, x_4, x_5, x_6, x_7, x_107); +if (lean_obj_tag(x_110) == 0) +{ +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; lean_object* x_126; +x_111 = lean_ctor_get(x_110, 1); +lean_inc(x_111); +lean_dec(x_110); +x_112 = lean_st_ref_take(x_3, x_111); +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); +lean_inc(x_114); +lean_dec(x_112); +x_115 = lean_ctor_get(x_113, 0); lean_inc(x_115); -x_116 = lean_ctor_get(x_114, 1); +x_116 = lean_ctor_get(x_113, 1); lean_inc(x_116); -lean_dec(x_114); -x_117 = lean_ctor_get(x_115, 0); +x_117 = lean_ctor_get(x_113, 2); lean_inc(x_117); -x_118 = lean_ctor_get(x_115, 1); +x_118 = lean_ctor_get(x_113, 3); lean_inc(x_118); -x_119 = lean_ctor_get(x_115, 2); +x_119 = lean_ctor_get(x_113, 4); lean_inc(x_119); -x_120 = lean_ctor_get(x_115, 3); -lean_inc(x_120); -x_121 = lean_ctor_get(x_115, 4); -lean_inc(x_121); -if (lean_is_exclusive(x_115)) { - lean_ctor_release(x_115, 0); - lean_ctor_release(x_115, 1); - lean_ctor_release(x_115, 2); - lean_ctor_release(x_115, 3); - lean_ctor_release(x_115, 4); - x_122 = x_115; -} else { - lean_dec_ref(x_115); - x_122 = lean_box(0); -} -x_123 = l_List_appendTR___rarg(x_119, x_14); -if (lean_is_scalar(x_122)) { - x_124 = lean_alloc_ctor(0, 5, 0); -} else { - x_124 = x_122; -} -lean_ctor_set(x_124, 0, x_117); -lean_ctor_set(x_124, 1, x_118); -lean_ctor_set(x_124, 2, x_123); -lean_ctor_set(x_124, 3, x_120); -lean_ctor_set(x_124, 4, x_121); -x_125 = lean_st_ref_set(x_5, x_124, x_116); -lean_dec(x_5); -x_126 = lean_ctor_get(x_125, 1); -lean_inc(x_126); -if (lean_is_exclusive(x_125)) { - lean_ctor_release(x_125, 0); - lean_ctor_release(x_125, 1); - x_127 = x_125; -} else { - lean_dec_ref(x_125); - x_127 = lean_box(0); -} -x_128 = lean_ctor_get(x_112, 0); -lean_inc(x_128); -lean_dec(x_112); -if (lean_is_scalar(x_127)) { - x_129 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_113)) { + lean_ctor_release(x_113, 0); + lean_ctor_release(x_113, 1); + lean_ctor_release(x_113, 2); + lean_ctor_release(x_113, 3); + lean_ctor_release(x_113, 4); + x_120 = x_113; } else { - x_129 = x_127; -} -lean_ctor_set(x_129, 0, x_128); -lean_ctor_set(x_129, 1, x_126); -return x_129; -} -block_148: -{ -lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; lean_object* x_137; lean_object* x_138; lean_object* x_139; lean_object* x_140; lean_object* x_141; lean_object* x_142; lean_object* x_143; lean_object* x_144; lean_object* x_145; lean_object* x_146; lean_object* x_147; -x_133 = lean_st_ref_take(x_5, x_132); -x_134 = lean_ctor_get(x_133, 0); -lean_inc(x_134); -x_135 = lean_ctor_get(x_133, 1); -lean_inc(x_135); -lean_dec(x_133); -x_136 = lean_ctor_get(x_134, 0); -lean_inc(x_136); -x_137 = lean_ctor_get(x_134, 1); -lean_inc(x_137); -x_138 = lean_ctor_get(x_134, 2); -lean_inc(x_138); -x_139 = lean_ctor_get(x_134, 3); -lean_inc(x_139); -x_140 = lean_ctor_get(x_134, 4); -lean_inc(x_140); -if (lean_is_exclusive(x_134)) { - lean_ctor_release(x_134, 0); - lean_ctor_release(x_134, 1); - lean_ctor_release(x_134, 2); - lean_ctor_release(x_134, 3); - lean_ctor_release(x_134, 4); - x_141 = x_134; -} else { - lean_dec_ref(x_134); - x_141 = lean_box(0); + lean_dec_ref(x_113); + x_120 = lean_box(0); } -x_142 = l_List_appendTR___rarg(x_138, x_14); -if (lean_is_scalar(x_141)) { - x_143 = lean_alloc_ctor(0, 5, 0); +x_121 = l_List_appendTR___rarg(x_117, x_12); +if (lean_is_scalar(x_120)) { + x_122 = lean_alloc_ctor(0, 5, 0); } else { - x_143 = x_141; + x_122 = x_120; } -lean_ctor_set(x_143, 0, x_136); -lean_ctor_set(x_143, 1, x_137); -lean_ctor_set(x_143, 2, x_142); -lean_ctor_set(x_143, 3, x_139); -lean_ctor_set(x_143, 4, x_140); -x_144 = lean_st_ref_set(x_5, x_143, x_135); -lean_dec(x_5); -x_145 = lean_ctor_get(x_144, 1); -lean_inc(x_145); -if (lean_is_exclusive(x_144)) { - lean_ctor_release(x_144, 0); - lean_ctor_release(x_144, 1); - x_146 = x_144; +lean_ctor_set(x_122, 0, x_115); +lean_ctor_set(x_122, 1, x_116); +lean_ctor_set(x_122, 2, x_121); +lean_ctor_set(x_122, 3, x_118); +lean_ctor_set(x_122, 4, x_119); +x_123 = lean_st_ref_set(x_3, x_122, x_114); +lean_dec(x_3); +x_124 = lean_ctor_get(x_123, 1); +lean_inc(x_124); +if (lean_is_exclusive(x_123)) { + lean_ctor_release(x_123, 0); + lean_ctor_release(x_123, 1); + x_125 = x_123; } else { - lean_dec_ref(x_144); - x_146 = lean_box(0); + lean_dec_ref(x_123); + x_125 = lean_box(0); } -if (lean_is_scalar(x_146)) { - x_147 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_125)) { + x_126 = lean_alloc_ctor(0, 2, 0); } else { - x_147 = x_146; - lean_ctor_set_tag(x_147, 1); -} -lean_ctor_set(x_147, 0, x_131); -lean_ctor_set(x_147, 1, x_145); -return x_147; -} -} + x_126 = x_125; } +lean_ctor_set(x_126, 0, x_106); +lean_ctor_set(x_126, 1, x_124); +return x_126; } -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp(lean_object* x_1) { -_start: +else { -lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 0); -return x_2; +lean_object* x_127; lean_object* x_128; +lean_dec(x_106); +x_127 = lean_ctor_get(x_110, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_110, 1); +lean_inc(x_128); +lean_dec(x_110); +x_87 = x_127; +x_88 = x_128; +goto block_104; } } -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___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) { -_start: +else { -lean_object* x_10; -x_10 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___lambda__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -lean_dec(x_8); +lean_object* x_129; lean_object* x_130; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_10; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { -_start: -{ -uint8_t x_11; uint8_t x_12; lean_object* x_13; -x_11 = lean_unbox(x_2); lean_dec(x_2); -x_12 = lean_unbox(x_3); -lean_dec(x_3); -x_13 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_1, x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_13; -} +x_129 = lean_ctor_get(x_105, 0); +lean_inc(x_129); +x_130 = lean_ctor_get(x_105, 1); +lean_inc(x_130); +lean_dec(x_105); +x_87 = x_129; +x_88 = x_130; +goto block_104; } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___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, lean_object* x_9, lean_object* x_10) { -_start: +block_104: { -uint8_t x_11; lean_object* x_12; -x_11 = 1; -x_12 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_3, x_1, x_11, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; -} +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_89 = lean_st_ref_take(x_3, x_88); +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_ctor_get(x_90, 0); +lean_inc(x_92); +x_93 = lean_ctor_get(x_90, 1); +lean_inc(x_93); +x_94 = lean_ctor_get(x_90, 2); +lean_inc(x_94); +x_95 = lean_ctor_get(x_90, 3); +lean_inc(x_95); +x_96 = lean_ctor_get(x_90, 4); +lean_inc(x_96); +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + lean_ctor_release(x_90, 2); + lean_ctor_release(x_90, 3); + lean_ctor_release(x_90, 4); + x_97 = x_90; +} else { + lean_dec_ref(x_90); + x_97 = lean_box(0); } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, uint8_t x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_box(x_4); -x_6 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withSynthesize___rarg___lambda__1___boxed), 10, 1); -lean_closure_set(x_6, 0, x_5); -x_7 = lean_apply_3(x_1, lean_box(0), x_6, x_3); -return x_7; +x_98 = l_List_appendTR___rarg(x_94, x_12); +if (lean_is_scalar(x_97)) { + x_99 = lean_alloc_ctor(0, 5, 0); +} else { + x_99 = x_97; +} +lean_ctor_set(x_99, 0, x_92); +lean_ctor_set(x_99, 1, x_93); +lean_ctor_set(x_99, 2, x_98); +lean_ctor_set(x_99, 3, x_95); +lean_ctor_set(x_99, 4, x_96); +x_100 = lean_st_ref_set(x_3, x_99, x_91); +lean_dec(x_3); +x_101 = lean_ctor_get(x_100, 1); +lean_inc(x_101); +if (lean_is_exclusive(x_100)) { + lean_ctor_release(x_100, 0); + lean_ctor_release(x_100, 1); + x_102 = x_100; +} else { + lean_dec_ref(x_100); + x_102 = lean_box(0); } +if (lean_is_scalar(x_102)) { + x_103 = lean_alloc_ctor(1, 2, 0); +} else { + x_103 = x_102; + lean_ctor_set_tag(x_103, 1); } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize(lean_object* x_1, lean_object* x_2) { -_start: -{ -lean_object* x_3; -x_3 = lean_alloc_closure((void*)(l_Lean_Elab_Term_withSynthesize___rarg___boxed), 4, 0); -return x_3; +lean_ctor_set(x_103, 0, x_87); +lean_ctor_set(x_103, 1, x_101); +return x_103; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___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) { -_start: -{ -uint8_t x_11; lean_object* x_12; -x_11 = lean_unbox(x_1); -lean_dec(x_1); -x_12 = l_Lean_Elab_Term_withSynthesize___rarg___lambda__1(x_11, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_12; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesize___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { +LEAN_EXPORT lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp(lean_object* x_1) { _start: { -uint8_t x_5; lean_object* x_6; -x_5 = lean_unbox(x_4); -lean_dec(x_4); -x_6 = l_Lean_Elab_Term_withSynthesize___rarg(x_1, x_2, x_3, x_5); -lean_dec(x_2); -return x_6; +lean_object* x_2; +x_2 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg), 8, 0); +return x_2; } } LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSynthesizeLight___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) { _start: { -uint8_t x_10; uint8_t x_11; lean_object* x_12; -x_10 = 1; -x_11 = 0; -x_12 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_2, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -return x_12; +lean_object* x_10; +x_10 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeLightImp___rarg(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +return x_10; } } static lean_object* _init_l_Lean_Elab_Term_withSynthesizeLight___rarg___closed__1() { @@ -20064,14 +21451,14 @@ x_16 = l_Lean_replaceRef(x_1, x_15); lean_dec(x_15); lean_dec(x_1); lean_ctor_set(x_7, 5, x_16); -x_17 = 0; +x_17 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_18 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, x_17, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_18 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_18) == 0) { lean_object* x_19; lean_object* x_20; lean_object* x_21; @@ -20163,14 +21550,14 @@ lean_ctor_set(x_40, 9, x_35); lean_ctor_set(x_40, 10, x_36); lean_ctor_set_uint8(x_40, sizeof(void*)*11, x_37); lean_ctor_set_uint8(x_40, sizeof(void*)*11 + 1, x_38); -x_41 = 0; +x_41 = 1; lean_inc(x_8); lean_inc(x_40); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_42 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, x_41, x_10, x_3, x_4, x_5, x_6, x_40, x_8, x_9); +x_42 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_13, x_41, x_3, x_4, x_5, x_6, x_40, x_8, x_9); if (lean_obj_tag(x_42) == 0) { lean_object* x_43; lean_object* x_44; lean_object* x_45; @@ -20514,7 +21901,7 @@ lean_dec(x_1); return x_14; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__1() { _start: { lean_object* x_1; @@ -20522,17 +21909,17 @@ x_1 = lean_mk_string_from_bytes("resume", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__2() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__1; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__3() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3() { _start: { lean_object* x_1; @@ -20540,27 +21927,27 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__4() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__5() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__4; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4; x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__6() { _start: { lean_object* x_1; @@ -20568,17 +21955,17 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__7() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__5; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__5; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__8() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__8() { _start: { lean_object* x_1; @@ -20586,17 +21973,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_SyntheticMVars___hyg_6210____closed__9() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__7; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__8; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__7; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__10() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10() { _start: { lean_object* x_1; @@ -20604,37 +21991,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__11() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__9; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__10; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__9; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__11; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__11; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__13() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__12; x_2 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_synthesizeUsingDefaultPrio_synthesizeUsingDefaultInstance___lambda__3___closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14() { _start: { lean_object* x_1; @@ -20642,17 +22029,17 @@ x_1 = lean_mk_string_from_bytes("SyntheticMVars", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____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_SyntheticMVars___hyg_6210____closed__13; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14; 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_SyntheticMVars___hyg_6210____closed__16() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16() { _start: { lean_object* x_1; @@ -20660,33 +22047,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_SyntheticMVars___hyg_6210____closed__17() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15; -x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__16; +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__15; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16; 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_SyntheticMVars___hyg_6210____closed__18() { +static lean_object* _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__17; -x_2 = lean_unsigned_to_nat(6210u); +x_1 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17; +x_2 = lean_unsigned_to_nat(6641u); 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_SyntheticMVars___hyg_6210_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641_(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_SyntheticMVars___hyg_6210____closed__2; +x_2 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2; x_3 = 0; -x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__18; +x_4 = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); return x_5; } @@ -20794,6 +22181,57 @@ l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___closed__2); l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___closed__3 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___closed__3(); lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_throwStuckAtUniverseCnstr___closed__3); +l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1 = _init_l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_PostponeBehavior_noConfusion___rarg___closed__1); +l_Lean_Elab_Term_instInhabitedPostponeBehavior = _init_l_Lean_Elab_Term_instInhabitedPostponeBehavior(); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__1 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__1(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__1); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__2); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__3); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__4); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__5); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__6); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__7); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__8); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__9); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__10); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__11); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__12); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__13); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__14); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__15); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__16); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__17); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__18); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__19); +l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20 = _init_l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20(); +lean_mark_persistent(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_reprPostponeBehavior____x40_Lean_Elab_SyntheticMVars___hyg_3986____closed__20); +l_Lean_Elab_Term_instReprPostponeBehavior___closed__1 = _init_l_Lean_Elab_Term_instReprPostponeBehavior___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_instReprPostponeBehavior___closed__1); +l_Lean_Elab_Term_instReprPostponeBehavior = _init_l_Lean_Elab_Term_instReprPostponeBehavior(); +lean_mark_persistent(l_Lean_Elab_Term_instReprPostponeBehavior); +l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1 = _init_l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_instBEqPostponeBehavior___closed__1); +l_Lean_Elab_Term_instBEqPostponeBehavior = _init_l_Lean_Elab_Term_instBEqPostponeBehavior(); +lean_mark_persistent(l_Lean_Elab_Term_instBEqPostponeBehavior); l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1(); lean_mark_persistent(l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__1); l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2 = _init_l_Lean_throwMaxRecDepthAt___at_Lean_Elab_Term_synthesizeSyntheticMVars_loop___spec__1___closed__2(); @@ -20892,43 +22330,43 @@ l_Lean_Elab_Term_withSynthesizeLight___rarg___closed__1 = _init_l_Lean_Elab_Term lean_mark_persistent(l_Lean_Elab_Term_withSynthesizeLight___rarg___closed__1); l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___spec__1___closed__1 = _init_l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___spec__1___closed__1(); lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_Elab_Term_runPendingTacticsAt___spec__1___closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__1); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__2(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__2); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__3); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__4(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__4); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__5(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__5); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__6); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__7(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__7); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__8(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__8); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__9(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__9); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__10(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__10); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__11(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__11); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__12); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__13(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__13); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__14); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__15); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__16 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__16(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__16); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__17 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__17(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__17); -l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__18 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__18(); -lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210____closed__18); -if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6210_(lean_io_mk_world()); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__1 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__1(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__1); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__2); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__3); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__4); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__5 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__5(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__5); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__6 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__6(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__6); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__7 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__7(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__7); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__8 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__8(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__8); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__9 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__9(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__9); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__10); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__11 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__11(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__11); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__12 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__12(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__12); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__13); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__14); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__15 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__15(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__15); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__16); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__17); +l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18 = _init_l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18(); +lean_mark_persistent(l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641____closed__18); +if (builtin) {res = l_Lean_Elab_Term_initFn____x40_Lean_Elab_SyntheticMVars___hyg_6641_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c index d566549140f8..ea5804ca2cdc 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c +++ b/stage0/stdlib/Lean/Elab/Tactic/BuiltinTactic.c @@ -824,7 +824,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalOpen(lean_object*, lean_object*, static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalUnknown_declRange___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_OpenDecl_resolveId___at_Lean_Elab_Tactic_evalOpen___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*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalIntros_declRange___closed__2; -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_elabSetOption___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_evalOpen___spec__44___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_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_evalAllGoals_declRange(lean_object*); @@ -20053,14 +20053,14 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalIntro_introStep___lambda__6(lean lean_object* x_12; uint8_t x_13; lean_object* x_14; x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Term_elabType), 8, 1); lean_closure_set(x_12, 0, x_1); -x_13 = 1; +x_13 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); if (lean_obj_tag(x_14) == 0) { lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; diff --git a/stage0/stdlib/Lean/Elab/Tactic/Conv/Rewrite.c b/stage0/stdlib/Lean/Elab/Tactic/Conv/Rewrite.c index 56bb6d957e3a..d93df14dcdc4 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Conv/Rewrite.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Conv/Rewrite.c @@ -40,7 +40,7 @@ static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRewrite___closed__7 lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalRewrite___lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRewrite_declRange___closed__4; -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRewrite_declRange___closed__7; static lean_object* l___regBuiltin_Lean_Elab_Tactic_Conv_evalRewrite___closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalRewrite___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -366,7 +366,7 @@ return x_60; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalRewrite___lambda__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; x_13 = lean_box(0); x_14 = lean_box(x_2); x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Conv_evalRewrite___lambda__1___boxed), 13, 4); @@ -378,10 +378,9 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_withMainContext___rarg), 10 lean_closure_set(x_16, 0, x_15); lean_closure_set(x_16, 1, x_4); lean_closure_set(x_16, 2, x_5); -x_17 = 0; -x_18 = 1; -x_19 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_16, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; +x_17 = 1; +x_18 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_18; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Conv_evalRewrite(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c index a84ac7766bf0..2cdccbc1ba12 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c +++ b/stage0/stdlib/Lean/Elab/Tactic/ElabTerm.c @@ -434,6 +434,7 @@ LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_evalConstructor_declRan lean_object* l_Lean_throwError___at_Lean_Elab_Tactic_evalTactic___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_evalDecide___rarg___lambda__2___closed__7; static lean_object* l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_preprocessPropToDecide___closed__1; +uint8_t l_Lean_Elab_Term_PostponeBehavior_ofBool(uint8_t); static lean_object* l_Lean_Elab_Tactic_getFVarId___closed__3; LEAN_EXPORT lean_object* l___private_Lean_Elab_Tactic_ElabTerm_0__Lean_Elab_Tactic_preprocessPropToDecide___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Tactic_evalRename___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*); @@ -516,88 +517,89 @@ 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) { -lean_object* x_11; lean_object* x_12; uint8_t x_13; lean_object* x_14; +lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; 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 = 0; -x_14 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_2, x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -if (lean_obj_tag(x_14) == 0) +x_13 = l_Lean_Elab_Term_PostponeBehavior_ofBool(x_2); +x_14 = 0; +x_15 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_13, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_12); +if (lean_obj_tag(x_15) == 0) { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_16; +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) { -lean_object* x_16; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -lean_ctor_set(x_14, 0, x_11); -return x_14; +lean_object* x_17; +x_17 = lean_ctor_get(x_15, 0); +lean_dec(x_17); +lean_ctor_set(x_15, 0, x_11); +return x_15; } else { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_17); -return x_18; +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_15, 1); +lean_inc(x_18); +lean_dec(x_15); +x_19 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_19, 0, x_11); +lean_ctor_set(x_19, 1, x_18); +return x_19; } } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_11); -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_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_23 = !lean_is_exclusive(x_10); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_10); +if (x_24 == 0) { return x_10; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_10, 0); -x_25 = lean_ctor_get(x_10, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +x_25 = lean_ctor_get(x_10, 0); +x_26 = lean_ctor_get(x_10, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); lean_dec(x_10); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_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; } } } @@ -6096,32 +6098,33 @@ lean_inc(x_8); x_16 = lean_apply_7(x_1, x_14, x_2, 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; uint8_t x_19; lean_object* x_20; +lean_object* x_17; lean_object* x_18; uint8_t x_19; uint8_t x_20; lean_object* x_21; x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); x_18 = lean_ctor_get(x_16, 1); lean_inc(x_18); lean_dec(x_16); -x_19 = 0; +x_19 = 1; +x_20 = 0; lean_inc(x_11); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -x_20 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_19, x_19, x_6, x_7, x_8, x_9, x_10, x_11, x_18); -if (lean_obj_tag(x_20) == 0) +x_21 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_19, x_20, x_6, x_7, x_8, x_9, x_10, x_11, x_18); +if (lean_obj_tag(x_21) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = l_Lean_Elab_Tactic_replaceMainGoal(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_21); -return x_22; +lean_object* x_22; lean_object* x_23; +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = l_Lean_Elab_Tactic_replaceMainGoal(x_17, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_11, x_22); +return x_23; } else { -uint8_t x_23; +uint8_t x_24; lean_dec(x_17); lean_dec(x_11); lean_dec(x_10); @@ -6131,29 +6134,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_20); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_21); +if (x_24 == 0) { -return x_20; +return x_21; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_20, 0); -x_25 = lean_ctor_get(x_20, 1); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_inc(x_24); -lean_dec(x_20); -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_21); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +return x_27; } } } else { -uint8_t x_27; +uint8_t x_28; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -6162,29 +6165,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_27 = !lean_is_exclusive(x_16); -if (x_27 == 0) +x_28 = !lean_is_exclusive(x_16); +if (x_28 == 0) { return x_16; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_16, 0); -x_29 = lean_ctor_get(x_16, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_16, 0); +x_30 = lean_ctor_get(x_16, 1); +lean_inc(x_30); lean_inc(x_29); -lean_inc(x_28); lean_dec(x_16); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +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_31; +uint8_t x_32; lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); @@ -6195,23 +6198,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_31 = !lean_is_exclusive(x_13); -if (x_31 == 0) +x_32 = !lean_is_exclusive(x_13); +if (x_32 == 0) { return x_13; } else { -lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_32 = lean_ctor_get(x_13, 0); -x_33 = lean_ctor_get(x_13, 1); +lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_33 = lean_ctor_get(x_13, 0); +x_34 = lean_ctor_get(x_13, 1); +lean_inc(x_34); lean_inc(x_33); -lean_inc(x_32); lean_dec(x_13); -x_34 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -return x_34; +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; } } } @@ -6254,34 +6257,35 @@ return x_21; } else { -uint8_t x_22; lean_object* x_23; -x_22 = 0; +uint8_t x_22; uint8_t x_23; lean_object* x_24; +x_22 = 1; +x_23 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_23 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_22, x_22, x_5, x_6, x_7, x_8, x_9, x_10, x_18); -if (lean_obj_tag(x_23) == 0) +x_24 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_22, x_23, x_5, x_6, x_7, x_8, x_9, x_10, x_18); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); -x_25 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__1(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 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; +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); +x_26 = l_Lean_instantiateMVars___at_Lean_Elab_Tactic_getMainTarget___spec__1(x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_box(0); -x_29 = l_Lean_Elab_Tactic_evalApplyLikeTactic___lambda__1(x_2, x_26, x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_27); -return x_29; +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +x_29 = lean_box(0); +x_30 = l_Lean_Elab_Tactic_evalApplyLikeTactic___lambda__1(x_2, x_27, x_29, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_28); +return x_30; } else { -uint8_t x_30; +uint8_t x_31; lean_dec(x_17); lean_dec(x_10); lean_dec(x_9); @@ -6292,30 +6296,30 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) +x_31 = !lean_is_exclusive(x_24); +if (x_31 == 0) { -return x_23; +return x_24; } 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_24, 0); +x_33 = lean_ctor_get(x_24, 1); +lean_inc(x_33); lean_inc(x_32); -lean_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_dec(x_24); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +return x_34; } } } } else { -uint8_t x_34; +uint8_t x_35; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -6325,23 +6329,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = !lean_is_exclusive(x_13); -if (x_34 == 0) +x_35 = !lean_is_exclusive(x_13); +if (x_35 == 0) { return x_13; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_13, 0); -x_36 = lean_ctor_get(x_13, 1); +lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_13, 0); +x_37 = lean_ctor_get(x_13, 1); +lean_inc(x_37); lean_inc(x_36); -lean_inc(x_35); lean_dec(x_13); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +x_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; } } } @@ -6623,32 +6627,33 @@ lean_inc(x_5); x_14 = l_Lean_MVarId_constructor(x_11, x_13, x_5, x_6, x_7, x_8, x_12); if (lean_obj_tag(x_14) == 0) { -lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; +lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); lean_dec(x_14); -x_17 = 0; +x_17 = 1; +x_18 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_18 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_17, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_16); -if (lean_obj_tag(x_18) == 0) +x_19 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_17, x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_16); +if (lean_obj_tag(x_19) == 0) { -lean_object* x_19; lean_object* x_20; -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); -x_20 = l_Lean_Elab_Tactic_replaceMainGoal(x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_19); -return x_20; +lean_object* x_20; lean_object* x_21; +x_20 = lean_ctor_get(x_19, 1); +lean_inc(x_20); +lean_dec(x_19); +x_21 = l_Lean_Elab_Tactic_replaceMainGoal(x_15, x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_20); +return x_21; } else { -uint8_t x_21; +uint8_t x_22; lean_dec(x_15); lean_dec(x_8); lean_dec(x_7); @@ -6658,29 +6663,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_21 = !lean_is_exclusive(x_18); -if (x_21 == 0) +x_22 = !lean_is_exclusive(x_19); +if (x_22 == 0) { -return x_18; +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 0); -x_23 = lean_ctor_get(x_18, 1); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 0); +x_24 = lean_ctor_get(x_19, 1); +lean_inc(x_24); lean_inc(x_23); -lean_inc(x_22); -lean_dec(x_18); -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_19); +x_25 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_25, 0, x_23); +lean_ctor_set(x_25, 1, x_24); +return x_25; } } } else { -uint8_t x_25; +uint8_t x_26; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -6689,29 +6694,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_25 = !lean_is_exclusive(x_14); -if (x_25 == 0) +x_26 = !lean_is_exclusive(x_14); +if (x_26 == 0) { return x_14; } else { -lean_object* x_26; lean_object* x_27; lean_object* x_28; -x_26 = lean_ctor_get(x_14, 0); -x_27 = lean_ctor_get(x_14, 1); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_14, 0); +x_28 = lean_ctor_get(x_14, 1); +lean_inc(x_28); lean_inc(x_27); -lean_inc(x_26); lean_dec(x_14); -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; +x_29 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_29, 0, x_27); +lean_ctor_set(x_29, 1, x_28); +return x_29; } } } else { -uint8_t x_29; +uint8_t x_30; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -6720,23 +6725,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_29 = !lean_is_exclusive(x_10); -if (x_29 == 0) +x_30 = !lean_is_exclusive(x_10); +if (x_30 == 0) { return x_10; } else { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_30 = lean_ctor_get(x_10, 0); -x_31 = lean_ctor_get(x_10, 1); +lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_31 = lean_ctor_get(x_10, 0); +x_32 = lean_ctor_get(x_10, 1); +lean_inc(x_32); lean_inc(x_31); -lean_inc(x_30); lean_dec(x_10); -x_32 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_32, 0, x_30); -lean_ctor_set(x_32, 1, x_31); -return x_32; +x_33 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_33, 0, x_31); +lean_ctor_set(x_33, 1, x_32); +return x_33; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Guard.c b/stage0/stdlib/Lean/Elab/Tactic/Guard.c index 46c2dbc01d84..418eebce9e8d 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Guard.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Guard.c @@ -1385,44 +1385,45 @@ lean_inc(x_7); x_25 = l_Lean_Meta_isExprDefEqGuarded(x_20, x_23, x_7, x_8, x_9, x_10, x_24); if (lean_obj_tag(x_25) == 0) { -lean_object* x_26; uint8_t x_27; lean_object* x_28; +lean_object* x_26; uint8_t x_27; uint8_t x_28; lean_object* x_29; x_26 = lean_ctor_get(x_25, 1); lean_inc(x_26); lean_dec(x_25); -x_27 = 0; +x_27 = 1; +x_28 = 0; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_28 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_27, x_27, x_5, x_6, x_7, x_8, x_9, x_10, x_26); -if (lean_obj_tag(x_28) == 0) +x_29 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_27, x_28, x_5, x_6, x_7, x_8, x_9, x_10, x_26); +if (lean_obj_tag(x_29) == 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 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); -x_30 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_29); -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 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; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); +x_31 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_30); +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); -x_33 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_32); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); +x_34 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_17, x_5, x_6, x_7, x_8, x_9, x_10, x_33); lean_dec(x_6); lean_dec(x_5); -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Elab_Tactic_GuardExpr_MatchKind_isEq(x_31, x_34, x_4, x_7, x_8, x_9, x_10, x_35); -return x_36; +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Elab_Tactic_GuardExpr_MatchKind_isEq(x_32, x_35, x_4, x_7, x_8, x_9, x_10, x_36); +return x_37; } else { -uint8_t x_37; +uint8_t x_38; lean_dec(x_17); lean_dec(x_14); lean_dec(x_10); @@ -1432,29 +1433,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_37 = !lean_is_exclusive(x_28); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_29); +if (x_38 == 0) { -return x_28; +return x_29; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_28, 0); -x_39 = lean_ctor_get(x_28, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_29, 0); +x_40 = lean_ctor_get(x_29, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_28); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_29); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } else { -uint8_t x_41; +uint8_t x_42; lean_dec(x_17); lean_dec(x_14); lean_dec(x_10); @@ -1464,29 +1465,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_41 = !lean_is_exclusive(x_25); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_25); +if (x_42 == 0) { return x_25; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_25, 0); -x_43 = lean_ctor_get(x_25, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_25, 0); +x_44 = lean_ctor_get(x_25, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); lean_dec(x_25); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_45; +uint8_t x_46; lean_dec(x_20); lean_dec(x_17); lean_dec(x_14); @@ -1497,29 +1498,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_45 = !lean_is_exclusive(x_22); -if (x_45 == 0) +x_46 = !lean_is_exclusive(x_22); +if (x_46 == 0) { return x_22; } else { -lean_object* x_46; lean_object* x_47; lean_object* x_48; -x_46 = lean_ctor_get(x_22, 0); -x_47 = lean_ctor_get(x_22, 1); +lean_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_22, 0); +x_48 = lean_ctor_get(x_22, 1); +lean_inc(x_48); lean_inc(x_47); -lean_inc(x_46); lean_dec(x_22); -x_48 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_48, 0, x_46); -lean_ctor_set(x_48, 1, x_47); -return x_48; +x_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_49; +uint8_t x_50; lean_dec(x_17); lean_dec(x_14); lean_dec(x_10); @@ -1529,29 +1530,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_49 = !lean_is_exclusive(x_19); -if (x_49 == 0) +x_50 = !lean_is_exclusive(x_19); +if (x_50 == 0) { return x_19; } else { -lean_object* x_50; lean_object* x_51; lean_object* x_52; -x_50 = lean_ctor_get(x_19, 0); -x_51 = lean_ctor_get(x_19, 1); +lean_object* x_51; lean_object* x_52; lean_object* x_53; +x_51 = lean_ctor_get(x_19, 0); +x_52 = lean_ctor_get(x_19, 1); +lean_inc(x_52); lean_inc(x_51); -lean_inc(x_50); lean_dec(x_19); -x_52 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_51); -return x_52; +x_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 { -uint8_t x_53; +uint8_t x_54; lean_dec(x_14); lean_dec(x_10); lean_dec(x_9); @@ -1560,29 +1561,29 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_53 = !lean_is_exclusive(x_16); -if (x_53 == 0) +x_54 = !lean_is_exclusive(x_16); +if (x_54 == 0) { return x_16; } else { -lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_54 = lean_ctor_get(x_16, 0); -x_55 = lean_ctor_get(x_16, 1); +lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_55 = lean_ctor_get(x_16, 0); +x_56 = lean_ctor_get(x_16, 1); +lean_inc(x_56); lean_inc(x_55); -lean_inc(x_54); lean_dec(x_16); -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; +x_57 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_57, 0, x_55); +lean_ctor_set(x_57, 1, x_56); +return x_57; } } } else { -uint8_t x_57; +uint8_t x_58; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -1592,23 +1593,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_57 = !lean_is_exclusive(x_13); -if (x_57 == 0) +x_58 = !lean_is_exclusive(x_13); +if (x_58 == 0) { return x_13; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_13, 0); -x_59 = lean_ctor_get(x_13, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_13, 0); +x_60 = lean_ctor_get(x_13, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); lean_dec(x_13); -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; +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; } } } @@ -5600,204 +5601,205 @@ lean_inc(x_3); x_12 = l_Lean_Elab_Term_elabTermEnsuringType(x_1, x_2, x_11, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(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; uint8_t x_24; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); -x_19 = lean_ctor_get(x_18, 0); -lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; uint8_t x_25; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); +x_20 = lean_ctor_get(x_19, 0); lean_inc(x_20); -lean_dec(x_18); -lean_inc(x_19); -x_21 = l_Lean_Meta_getMVars(x_19, x_6, x_7, x_8, x_9, x_20); -x_22 = lean_ctor_get(x_21, 0); -lean_inc(x_22); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = l_Array_isEmpty___rarg(x_22); -if (x_24 == 0) -{ -lean_object* x_25; +x_21 = lean_ctor_get(x_19, 1); +lean_inc(x_21); lean_dec(x_19); -x_25 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_23); +lean_inc(x_20); +x_22 = l_Lean_Meta_getMVars(x_20, x_6, x_7, x_8, x_9, x_21); +x_23 = lean_ctor_get(x_22, 0); +lean_inc(x_23); +x_24 = lean_ctor_get(x_22, 1); +lean_inc(x_24); lean_dec(x_22); -if (lean_obj_tag(x_25) == 0) +x_25 = l_Array_isEmpty___rarg(x_23); +if (x_25 == 0) { -uint8_t x_26; -x_26 = !lean_is_exclusive(x_25); -if (x_26 == 0) +lean_object* x_26; +lean_dec(x_20); +x_26 = l_Lean_Elab_Term_logUnassignedUsingErrorInfos(x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_24); +lean_dec(x_23); +if (lean_obj_tag(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; +uint8_t x_27; +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_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_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 { -uint8_t x_32; -x_32 = !lean_is_exclusive(x_25); -if (x_32 == 0) +uint8_t x_33; +x_33 = !lean_is_exclusive(x_26); +if (x_33 == 0) { -return x_25; +return x_26; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_25, 0); -x_34 = lean_ctor_get(x_25, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_26, 0); +x_35 = lean_ctor_get(x_26, 1); +lean_inc(x_35); lean_inc(x_34); -lean_inc(x_33); -lean_dec(x_25); -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_dec(x_26); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } else { -lean_object* x_36; -lean_dec(x_22); +lean_object* x_37; +lean_dec(x_23); lean_dec(x_3); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); -lean_inc(x_19); -x_36 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd_unsafe__1(x_19, x_6, x_7, x_8, x_9, x_23); -if (lean_obj_tag(x_36) == 0) +lean_inc(x_20); +x_37 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd_unsafe__1(x_20, x_6, x_7, x_8, x_9, x_24); +if (lean_obj_tag(x_37) == 0) { -lean_object* x_37; uint8_t x_38; -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_unbox(x_37); -lean_dec(x_37); -if (x_38 == 0) +lean_object* x_38; uint8_t x_39; +x_38 = lean_ctor_get(x_37, 0); +lean_inc(x_38); +x_39 = lean_unbox(x_38); +lean_dec(x_38); +if (x_39 == 0) { -lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_39 = lean_ctor_get(x_36, 1); -lean_inc(x_39); -lean_dec(x_36); -x_40 = l_Lean_indentExpr(x_19); -x_41 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd___lambda__1___closed__2; -x_42 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_40); -x_43 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd___lambda__1___closed__4; -x_44 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -x_45 = l_Lean_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_39); +lean_object* x_40; 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_ctor_get(x_37, 1); +lean_inc(x_40); +lean_dec(x_37); +x_41 = l_Lean_indentExpr(x_20); +x_42 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd___lambda__1___closed__2; +x_43 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_43, 0, x_42); +lean_ctor_set(x_43, 1, x_41); +x_44 = l_Lean_Elab_Tactic_GuardExpr_evalGuardCmd___lambda__1___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_throwError___at_Lean_Elab_Term_throwErrorIfErrors___spec__1(x_45, x_4, x_5, x_6, x_7, x_8, x_9, x_40); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -return x_45; +return x_46; } else { -uint8_t x_46; -lean_dec(x_19); +uint8_t x_47; +lean_dec(x_20); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_46 = !lean_is_exclusive(x_36); -if (x_46 == 0) -{ -lean_object* x_47; lean_object* x_48; -x_47 = lean_ctor_get(x_36, 0); -lean_dec(x_47); -x_48 = lean_box(0); -lean_ctor_set(x_36, 0, x_48); -return x_36; +x_47 = !lean_is_exclusive(x_37); +if (x_47 == 0) +{ +lean_object* x_48; lean_object* x_49; +x_48 = lean_ctor_get(x_37, 0); +lean_dec(x_48); +x_49 = lean_box(0); +lean_ctor_set(x_37, 0, x_49); +return x_37; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_36, 1); -lean_inc(x_49); -lean_dec(x_36); -x_50 = lean_box(0); -x_51 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_51, 0, x_50); -lean_ctor_set(x_51, 1, x_49); -return x_51; +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_37, 1); +lean_inc(x_50); +lean_dec(x_37); +x_51 = lean_box(0); +x_52 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_52, 0, x_51); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } } else { -uint8_t x_52; -lean_dec(x_19); +uint8_t x_53; +lean_dec(x_20); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_52 = !lean_is_exclusive(x_36); -if (x_52 == 0) +x_53 = !lean_is_exclusive(x_37); +if (x_53 == 0) { -return x_36; +return x_37; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_36, 0); -x_54 = lean_ctor_get(x_36, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_37, 0); +x_55 = lean_ctor_get(x_37, 1); +lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_36); -x_55 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_55, 0, x_53); -lean_ctor_set(x_55, 1, x_54); -return x_55; +lean_dec(x_37); +x_56 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_56, 0, x_54); +lean_ctor_set(x_56, 1, x_55); +return x_56; } } } } else { -uint8_t x_56; +uint8_t x_57; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -5806,29 +5808,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_56 = !lean_is_exclusive(x_16); -if (x_56 == 0) +x_57 = !lean_is_exclusive(x_17); +if (x_57 == 0) { -return x_16; +return x_17; } else { -lean_object* x_57; lean_object* x_58; lean_object* x_59; -x_57 = lean_ctor_get(x_16, 0); -x_58 = lean_ctor_get(x_16, 1); +lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_58 = lean_ctor_get(x_17, 0); +x_59 = lean_ctor_get(x_17, 1); +lean_inc(x_59); lean_inc(x_58); -lean_inc(x_57); -lean_dec(x_16); -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_dec(x_17); +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_60; +uint8_t x_61; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -5836,23 +5838,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_60 = !lean_is_exclusive(x_12); -if (x_60 == 0) +x_61 = !lean_is_exclusive(x_12); +if (x_61 == 0) { return x_12; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; -x_61 = lean_ctor_get(x_12, 0); -x_62 = lean_ctor_get(x_12, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; +x_62 = lean_ctor_get(x_12, 0); +x_63 = lean_ctor_get(x_12, 1); +lean_inc(x_63); lean_inc(x_62); -lean_inc(x_61); lean_dec(x_12); -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; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Induction.c b/stage0/stdlib/Lean/Elab/Tactic/Induction.c index eab8a542618a..8d545d876539 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Induction.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Induction.c @@ -15643,142 +15643,143 @@ lean_inc(x_3); x_12 = l_Lean_Elab_Term_elabTerm(x_1, x_2, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; x_13 = lean_ctor_get(x_12, 0); lean_inc(x_13); x_14 = lean_ctor_get(x_12, 1); lean_inc(x_14); lean_dec(x_12); +x_15 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_15 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_11, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_14); -if (lean_obj_tag(x_15) == 0) +x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +if (lean_obj_tag(x_16) == 0) { -lean_object* x_16; lean_object* x_17; uint8_t x_18; -x_16 = lean_ctor_get(x_15, 1); -lean_inc(x_16); -lean_dec(x_15); -x_17 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_16); +lean_object* x_17; lean_object* x_18; uint8_t x_19; +x_17 = lean_ctor_get(x_16, 1); +lean_inc(x_17); +lean_dec(x_16); +x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_3, x_4, x_5, x_6, x_7, x_8, x_17); lean_dec(x_4); lean_dec(x_3); -x_18 = !lean_is_exclusive(x_17); -if (x_18 == 0) +x_19 = !lean_is_exclusive(x_18); +if (x_19 == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_19 = lean_ctor_get(x_17, 0); -x_20 = lean_ctor_get(x_17, 1); -x_21 = l_Lean_Expr_eta(x_19); -x_22 = l_Lean_Expr_hasMVar(x_21); -if (x_22 == 0) +lean_object* x_20; lean_object* x_21; lean_object* x_22; uint8_t x_23; +x_20 = lean_ctor_get(x_18, 0); +x_21 = lean_ctor_get(x_18, 1); +x_22 = l_Lean_Expr_eta(x_20); +x_23 = l_Lean_Expr_hasMVar(x_22); +if (x_23 == 0) { lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -lean_ctor_set(x_17, 0, x_21); -return x_17; +lean_ctor_set(x_18, 0, x_22); +return x_18; } else { -lean_object* x_23; uint8_t x_24; -lean_free_object(x_17); -x_23 = l_Lean_Meta_abstractMVars(x_21, x_11, x_5, x_6, x_7, x_8, x_20); +lean_object* x_24; uint8_t x_25; +lean_free_object(x_18); +x_24 = l_Lean_Meta_abstractMVars(x_22, x_11, x_5, x_6, x_7, x_8, x_21); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_24 = !lean_is_exclusive(x_23); -if (x_24 == 0) +x_25 = !lean_is_exclusive(x_24); +if (x_25 == 0) { -lean_object* x_25; lean_object* x_26; -x_25 = lean_ctor_get(x_23, 0); -x_26 = lean_ctor_get(x_25, 2); -lean_inc(x_26); -lean_dec(x_25); -lean_ctor_set(x_23, 0, x_26); -return x_23; +lean_object* x_26; lean_object* x_27; +x_26 = lean_ctor_get(x_24, 0); +x_27 = lean_ctor_get(x_26, 2); +lean_inc(x_27); +lean_dec(x_26); +lean_ctor_set(x_24, 0, x_27); +return x_24; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_27 = lean_ctor_get(x_23, 0); -x_28 = lean_ctor_get(x_23, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_23); -x_29 = lean_ctor_get(x_27, 2); +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_dec(x_27); -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_inc(x_28); +lean_dec(x_24); +x_30 = lean_ctor_get(x_28, 2); +lean_inc(x_30); +lean_dec(x_28); +x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_31, 0, x_30); +lean_ctor_set(x_31, 1, x_29); +return x_31; } } } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; uint8_t x_34; -x_31 = lean_ctor_get(x_17, 0); -x_32 = lean_ctor_get(x_17, 1); +lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_32 = lean_ctor_get(x_18, 0); +x_33 = lean_ctor_get(x_18, 1); +lean_inc(x_33); lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_17); -x_33 = l_Lean_Expr_eta(x_31); -x_34 = l_Lean_Expr_hasMVar(x_33); -if (x_34 == 0) +lean_dec(x_18); +x_34 = l_Lean_Expr_eta(x_32); +x_35 = l_Lean_Expr_hasMVar(x_34); +if (x_35 == 0) { -lean_object* x_35; +lean_object* x_36; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_33); -lean_ctor_set(x_35, 1, x_32); -return x_35; +x_36 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_33); +return x_36; } 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; -x_36 = l_Lean_Meta_abstractMVars(x_33, x_11, x_5, x_6, x_7, x_8, x_32); +lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; lean_object* x_42; +x_37 = l_Lean_Meta_abstractMVars(x_34, x_11, x_5, x_6, x_7, x_8, x_33); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); +x_38 = lean_ctor_get(x_37, 0); lean_inc(x_38); -if (lean_is_exclusive(x_36)) { - lean_ctor_release(x_36, 0); - lean_ctor_release(x_36, 1); - x_39 = x_36; +x_39 = lean_ctor_get(x_37, 1); +lean_inc(x_39); +if (lean_is_exclusive(x_37)) { + lean_ctor_release(x_37, 0); + lean_ctor_release(x_37, 1); + x_40 = x_37; } else { - lean_dec_ref(x_36); - x_39 = lean_box(0); + lean_dec_ref(x_37); + x_40 = lean_box(0); } -x_40 = lean_ctor_get(x_37, 2); -lean_inc(x_40); -lean_dec(x_37); -if (lean_is_scalar(x_39)) { - x_41 = lean_alloc_ctor(0, 2, 0); +x_41 = lean_ctor_get(x_38, 2); +lean_inc(x_41); +lean_dec(x_38); +if (lean_is_scalar(x_40)) { + x_42 = lean_alloc_ctor(0, 2, 0); } else { - x_41 = x_39; + x_42 = x_40; } -lean_ctor_set(x_41, 0, x_40); -lean_ctor_set(x_41, 1, x_38); -return x_41; +lean_ctor_set(x_42, 0, x_41); +lean_ctor_set(x_42, 1, x_39); +return x_42; } } } else { -uint8_t x_42; +uint8_t x_43; lean_dec(x_13); lean_dec(x_8); lean_dec(x_7); @@ -15786,52 +15787,52 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_42 = !lean_is_exclusive(x_15); -if (x_42 == 0) +x_43 = !lean_is_exclusive(x_16); +if (x_43 == 0) { -return x_15; +return x_16; } else { -lean_object* x_43; lean_object* x_44; lean_object* x_45; -x_43 = lean_ctor_get(x_15, 0); -x_44 = lean_ctor_get(x_15, 1); +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_inc(x_43); -lean_dec(x_15); -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_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; } } } else { -uint8_t x_46; +uint8_t x_47; 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_12); -if (x_46 == 0) +x_47 = !lean_is_exclusive(x_12); +if (x_47 == 0) { return x_12; } else { -lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_47 = lean_ctor_get(x_12, 0); -x_48 = lean_ctor_get(x_12, 1); +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_inc(x_48); -lean_inc(x_47); lean_dec(x_12); -x_49 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_49, 0, x_47); -lean_ctor_set(x_49, 1, x_48); -return x_49; +x_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; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/LibrarySearch.c b/stage0/stdlib/Lean/Elab/Tactic/LibrarySearch.c index f648a258a346..fa41a6e6932f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/LibrarySearch.c +++ b/stage0/stdlib/Lean/Elab/Tactic/LibrarySearch.c @@ -666,6 +666,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_LibrarySearch_exact_x3f___lambda__5(lean_ob lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; size_t x_21; lean_object* x_22; lean_object* x_23; x_16 = l_Lean_Elab_LibrarySearch_exact_x3f___lambda__5___closed__2; x_17 = lean_unsigned_to_nat(90u); +lean_inc(x_1); x_18 = l_Lean_reportOutOfHeartbeats(x_16, x_1, x_17, x_13, x_14, x_15); x_19 = lean_ctor_get(x_18, 1); lean_inc(x_19); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Meta.c b/stage0/stdlib/Lean/Elab/Tactic/Meta.c index 3e33ef41ed2e..7264148bfeda 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Meta.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Meta.c @@ -29,7 +29,7 @@ lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlFromMAux___at_Lean_Elab_runTactic___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, 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_Lean_instantiateMVarDeclMVars___at_Lean_Elab_runTactic___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___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_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_EXPORT lean_object* l___private_Lean_Data_PersistentArray_0__Lean_PersistentArray_foldlMAux___at_Lean_Elab_runTactic___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -1399,7 +1399,7 @@ return x_17; LEAN_EXPORT lean_object* l_Lean_Elab_runTactic(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); @@ -1414,16 +1414,13 @@ lean_closure_set(x_12, 0, x_2); x_13 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_run), 9, 2); lean_closure_set(x_13, 0, x_1); lean_closure_set(x_13, 1, x_12); -x_14 = 0; -x_15 = 1; -x_16 = lean_box(x_14); -x_17 = lean_box(x_15); -x_18 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_18, 0, x_13); -lean_closure_set(x_18, 1, x_16); -lean_closure_set(x_18, 2, x_17); -x_19 = l_Lean_Elab_Term_TermElabM_run___rarg(x_18, x_3, x_4, x_5, x_6, x_7, x_8, x_11); -return x_19; +x_14 = 1; +x_15 = lean_box(x_14); +x_16 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_16, 0, x_13); +lean_closure_set(x_16, 1, x_15); +x_17 = l_Lean_Elab_Term_TermElabM_run___rarg(x_16, x_3, x_4, x_5, x_6, x_7, x_8, x_11); +return x_17; } } LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_runTactic___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) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c index ef27063f241d..5cf43abcaef4 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/NormCast.c +++ b/stage0/stdlib/Lean/Elab/Tactic/NormCast.c @@ -13632,74 +13632,75 @@ lean_inc(x_4); x_17 = l_Lean_Elab_Term_elabTerm(x_2, x_15, x_16, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_13); if (lean_obj_tag(x_17) == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; +lean_object* x_18; lean_object* x_19; uint8_t x_20; uint8_t x_21; lean_object* x_22; 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 = 0; +x_21 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_21 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_16, x_20, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_21) == 0) +x_22 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_21, x_4, x_5, x_6, x_7, x_8, x_9, x_19); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_22; lean_object* x_23; -x_22 = lean_ctor_get(x_21, 1); -lean_inc(x_22); -lean_dec(x_21); +lean_object* x_23; lean_object* x_24; +x_23 = lean_ctor_get(x_22, 1); +lean_inc(x_23); +lean_dec(x_22); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_18); -x_23 = lean_infer_type(x_18, x_6, x_7, x_8, x_9, x_22); -if (lean_obj_tag(x_23) == 0) +x_24 = lean_infer_type(x_18, x_6, x_7, x_8, x_9, x_23); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; uint8_t x_29; -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; uint8_t x_30; +x_25 = lean_ctor_get(x_24, 0); lean_inc(x_25); -lean_dec(x_23); -x_26 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_24, x_4, x_5, x_6, x_7, x_8, x_9, x_25); -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); +x_26 = lean_ctor_get(x_24, 1); +lean_inc(x_26); +lean_dec(x_24); +x_27 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_25, x_4, x_5, x_6, x_7, x_8, x_9, x_26); +x_28 = lean_ctor_get(x_27, 0); lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Expr_hasExprMVar(x_27); -if (x_29 == 0) +x_29 = lean_ctor_get(x_27, 1); +lean_inc(x_29); +lean_dec(x_27); +x_30 = l_Lean_Expr_hasExprMVar(x_28); +if (x_30 == 0) { -lean_object* x_30; lean_object* x_31; -x_30 = lean_box(0); -x_31 = l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2(x_27, x_14, x_1, x_12, x_18, x_30, x_4, x_5, x_6, x_7, x_8, x_9, x_28); -return x_31; +lean_object* x_31; lean_object* x_32; +x_31 = lean_box(0); +x_32 = l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2(x_28, x_14, x_1, x_12, x_18, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_29); +return x_32; } else { -lean_object* x_32; +lean_object* x_33; lean_inc(x_4); -x_32 = l_Lean_Elab_Term_tryPostpone(x_4, x_5, x_6, x_7, x_8, x_9, x_28); -if (lean_obj_tag(x_32) == 0) +x_33 = l_Lean_Elab_Term_tryPostpone(x_4, x_5, x_6, x_7, x_8, x_9, x_29); +if (lean_obj_tag(x_33) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_33, 0); lean_inc(x_34); -lean_dec(x_32); -x_35 = l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2(x_27, x_14, x_1, x_12, x_18, x_33, x_4, x_5, x_6, x_7, x_8, x_9, x_34); -return x_35; +x_35 = lean_ctor_get(x_33, 1); +lean_inc(x_35); +lean_dec(x_33); +x_36 = l_Lean_Elab_Tactic_NormCast_elabModCast___lambda__2(x_28, x_14, x_1, x_12, x_18, x_34, x_4, x_5, x_6, x_7, x_8, x_9, x_35); +return x_36; } else { -uint8_t x_36; -lean_dec(x_27); +uint8_t x_37; +lean_dec(x_28); lean_dec(x_18); lean_dec(x_14); lean_dec(x_12); @@ -13710,30 +13711,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_36 = !lean_is_exclusive(x_32); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_33); +if (x_37 == 0) { -return x_32; +return x_33; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_32, 0); -x_38 = lean_ctor_get(x_32, 1); +lean_object* x_38; lean_object* x_39; lean_object* x_40; +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_inc(x_37); -lean_dec(x_32); -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_dec(x_33); +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 { -uint8_t x_40; +uint8_t x_41; lean_dec(x_18); lean_dec(x_14); lean_dec(x_12); @@ -13744,29 +13745,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_40 = !lean_is_exclusive(x_23); -if (x_40 == 0) +x_41 = !lean_is_exclusive(x_24); +if (x_41 == 0) { -return x_23; +return x_24; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_23, 0); -x_42 = lean_ctor_get(x_23, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_24, 0); +x_43 = lean_ctor_get(x_24, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_23); -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_dec(x_24); +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_44; +uint8_t x_45; lean_dec(x_18); lean_dec(x_14); lean_dec(x_12); @@ -13777,29 +13778,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_44 = !lean_is_exclusive(x_21); -if (x_44 == 0) +x_45 = !lean_is_exclusive(x_22); +if (x_45 == 0) { -return x_21; +return x_22; } else { -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_object* x_46; lean_object* x_47; lean_object* x_48; +x_46 = lean_ctor_get(x_22, 0); +x_47 = lean_ctor_get(x_22, 1); +lean_inc(x_47); 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; +lean_dec(x_22); +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_48; +uint8_t x_49; lean_dec(x_14); lean_dec(x_12); lean_dec(x_9); @@ -13809,29 +13810,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_48 = !lean_is_exclusive(x_17); -if (x_48 == 0) +x_49 = !lean_is_exclusive(x_17); +if (x_49 == 0) { return x_17; } else { -lean_object* x_49; lean_object* x_50; lean_object* x_51; -x_49 = lean_ctor_get(x_17, 0); -x_50 = lean_ctor_get(x_17, 1); +lean_object* x_50; lean_object* x_51; lean_object* x_52; +x_50 = lean_ctor_get(x_17, 0); +x_51 = lean_ctor_get(x_17, 1); +lean_inc(x_51); lean_inc(x_50); -lean_inc(x_49); lean_dec(x_17); -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; +x_52 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_52, 0, x_50); +lean_ctor_set(x_52, 1, x_51); +return x_52; } } } else { -uint8_t x_52; +uint8_t x_53; lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -13840,23 +13841,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_52 = !lean_is_exclusive(x_11); -if (x_52 == 0) +x_53 = !lean_is_exclusive(x_11); +if (x_53 == 0) { return x_11; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; -x_53 = lean_ctor_get(x_11, 0); -x_54 = lean_ctor_get(x_11, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; +x_54 = lean_ctor_get(x_11, 0); +x_55 = lean_ctor_get(x_11, 1); +lean_inc(x_55); lean_inc(x_54); -lean_inc(x_53); lean_dec(x_11); -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_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; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c index 853bc5f61180..2d050f1bdd4f 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Omega/Frontend.c @@ -34,11 +34,13 @@ 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_omegaImpl___lambda__4___closed__6; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___closed__5; +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1; 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*, 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*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__4; 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*); @@ -64,6 +66,7 @@ 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_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921_(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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_prettyCoeffs___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__4; @@ -86,7 +89,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___cl 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*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__20___closed__1; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_problem___default___closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleFinVal___closed__13; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__31; @@ -255,6 +257,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__6(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__22___closed__7; static lean_object* l_Lean_Elab_Tactic_Omega_omegaImpl___lambda__4___closed__4; +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__39; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__24; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__3; @@ -270,7 +273,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__4___clo static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_asLinearComboImpl___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_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_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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -333,7 +336,6 @@ static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_ 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*); -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922_(lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_formatErrorMessage_inScope___closed__1; uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__1; @@ -378,6 +380,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSa static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__26; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addFact___lambda__1___closed__12; LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapM___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__8(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__14; static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__4___closed__2; static lean_object* l_Std_Range_forIn_x27_loop___at_Lean_Elab_Tactic_Omega_formatErrorMessage_varNames___spec__2___closed__1; @@ -436,7 +439,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_addIntEquality___l 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; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3; 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*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -466,14 +468,12 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__ extern lean_object* l_Lean_levelZero; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_rewrite___lambda__2___closed__3; static lean_object* l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__5; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2; static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl_handleNatCast___closed__47; 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*, 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*); -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4; 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; static lean_object* l_Lean_Elab_Tactic_Omega_mkEvalRflProof___closed__3; @@ -489,6 +489,7 @@ static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___ 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_object*, lean_object*); +static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__3; 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_mkEvalRflProof___closed__6; @@ -496,7 +497,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_trivial; lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_MVarId_elabFalseOrByContra___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange(lean_object*); extern lean_object* l_String_instInhabited; -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; @@ -861,7 +861,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction(lean_object*, static lean_object* l_Lean_Elab_Tactic_Omega_asLinearComboImpl___lambda__8___closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_splitDisjunction___lambda__8___boxed(lean_object**); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__3; -static lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1; lean_object* l_Lean_Meta_evalExpr_x27___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_Omega_MetaProblem_pushNot___lambda__22___closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_Omega_MetaProblem_addIntInequality___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -3299,38 +3298,39 @@ lean_inc(x_4); x_12 = l_Lean_Elab_Term_elabTermEnsuringType(x_1, x_2, x_11, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_18; +return x_19; } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -3338,52 +3338,52 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -return x_16; +return x_17; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -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_17); +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_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_12); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) { return x_12; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -x_25 = lean_ctor_get(x_12, 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_12); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_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; } } } @@ -3515,7 +3515,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -3527,85 +3527,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_Omega_elabOmegaConfig___lam lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__7; -x_24 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__7; +x_22 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_Omega_elabOmegaConfig___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__11; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_Omega_elabOmegaConfig___closed__11; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } } } @@ -34511,17 +34508,6 @@ return x_43; } } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; -x_1 = 0; -x_2 = lean_box(x_1); -x_3 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_3, 0, x_2); -return x_3; -} -} LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -34543,7 +34529,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1; +x_14 = lean_box(0); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); @@ -35188,7 +35174,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRa _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(680u); x_2 = lean_unsigned_to_nat(0u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35200,7 +35186,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRa _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = lean_unsigned_to_nat(685u); +x_1 = lean_unsigned_to_nat(684u); x_2 = lean_unsigned_to_nat(31u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35228,7 +35214,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRa _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(680u); x_2 = lean_unsigned_to_nat(4u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35240,7 +35226,7 @@ static lean_object* _init_l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRa _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(680u); x_2 = lean_unsigned_to_nat(13u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -35286,7 +35272,7 @@ x_4 = l_Lean_addBuiltinDeclarationRanges(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1() { _start: { lean_object* x_1; @@ -35294,17 +35280,17 @@ x_1 = lean_mk_string_from_bytes("bv_toNat", 8); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1; +x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__3() { _start: { lean_object* x_1; @@ -35312,7 +35298,7 @@ x_1 = lean_mk_string_from_bytes("bvOmegaSimpExtension", 20); return x_1; } } -static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____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; @@ -35320,12 +35306,12 @@ x_1 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend x_2 = l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega___closed__1; x_3 = l_Lean_Elab_Tactic_Omega_evalOmega___closed__2; x_4 = l_Lean_Elab_Tactic_Omega_evalUnsafe____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_6____closed__3; -x_5 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3; +x_5 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____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_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5() { +static lean_object* _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5() { _start: { lean_object* x_1; @@ -35333,13 +35319,13 @@ x_1 = lean_mk_string_from_bytes("simp lemmas converting `BitVec` goals to `Nat` return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2; -x_3 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5; -x_4 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4; +x_2 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2; +x_3 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5; +x_4 = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__4; x_5 = l_Lean_Meta_registerSimpAttr(x_2, x_3, x_4, x_1); return x_5; } @@ -36303,8 +36289,6 @@ l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2___closed__3 = _init_l_Lean_Elab lean_mark_persistent(l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2___closed__3); l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2___closed__4 = _init_l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2___closed__4(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__2___closed__4); -l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1 = _init_l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_omegaTactic___lambda__3___closed__1); l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1___closed__1 = _init_l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1___closed__1(); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_evalOmega___lambda__1___closed__1); l_Lean_Elab_Tactic_Omega_evalOmega___closed__1 = _init_l_Lean_Elab_Tactic_Omega_evalOmega___closed__1(); @@ -36345,17 +36329,17 @@ lean_mark_persistent(l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange__ if (builtin) {res = l___regBuiltin_Lean_Elab_Tactic_Omega_evalOmega_declRange(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__1); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__2); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__3); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__4); -l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5(); -lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922____closed__5); -if (builtin) {res = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15922_(lean_io_mk_world()); +}l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__1); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__2); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__3 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__3(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__3); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__4 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__4(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__4); +l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5 = _init_l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5(); +lean_mark_persistent(l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921____closed__5); +if (builtin) {res = l_Lean_Elab_Tactic_Omega_initFn____x40_Lean_Elab_Tactic_Omega_Frontend___hyg_15921_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_Tactic_Omega_bvOmegaSimpExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_Tactic_Omega_bvOmegaSimpExtension); diff --git a/stage0/stdlib/Lean/Elab/Tactic/RCases.c b/stage0/stdlib/Lean/Elab/Tactic/RCases.c index e91384789f15..8d8829666ff4 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/RCases.c +++ b/stage0/stdlib/Lean/Elab/Tactic/RCases.c @@ -395,7 +395,7 @@ static lean_object* l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__11___c LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_rcasesCore___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at_Lean_Elab_Tactic_RCases_RCasesPatt_parse___spec__1___rarg___closed__1; static lean_object* l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_reprRCasesPatt____x40_Lean_Elab_Tactic_RCases___hyg_933____closed__3; -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_RCases_evalRIntro___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_expandRIntroPat___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Tactic_RCases_0__Lean_Elab_Tactic_RCases_reprRCasesPatt____x40_Lean_Elab_Tactic_RCases___hyg_933____closed__11; @@ -7135,105 +7135,106 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_rcasesCore___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_15; lean_object* x_16; -x_15 = 0; +uint8_t x_15; uint8_t x_16; lean_object* x_17; +x_15 = 1; +x_16 = 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_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_8, x_9, x_10, x_11, x_12, x_13, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_14); +if (lean_obj_tag(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; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_box(0); +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = lean_box(0); lean_inc(x_1); -x_19 = l_Lean_Expr_mdata___override(x_18, x_1); -x_20 = lean_box(0); +x_20 = l_Lean_Expr_mdata___override(x_19, x_1); x_21 = lean_box(0); +x_22 = 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_22 = l_Lean_Elab_Term_addTermInfo_x27(x_2, x_19, x_20, x_20, x_21, x_15, x_8, x_9, x_10, x_11, x_12, x_13, x_17); -if (lean_obj_tag(x_22) == 0) -{ -uint8_t x_23; -x_23 = !lean_is_exclusive(x_22); -if (x_23 == 0) +x_23 = l_Lean_Elab_Term_addTermInfo_x27(x_2, x_20, x_21, x_21, x_22, x_16, x_8, x_9, x_10, x_11, x_12, x_13, x_18); +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_22, 1); -x_25 = lean_ctor_get(x_22, 0); -lean_dec(x_25); -x_26 = l_Lean_Meta_FVarSubst_apply(x_3, x_1); -if (lean_obj_tag(x_26) == 1) +uint8_t x_24; +x_24 = !lean_is_exclusive(x_23); +if (x_24 == 0) { -lean_object* x_45; -x_45 = lean_ctor_get(x_26, 0); -lean_inc(x_45); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_25 = lean_ctor_get(x_23, 1); +x_26 = lean_ctor_get(x_23, 0); lean_dec(x_26); -lean_ctor_set(x_22, 0, x_45); -x_27 = x_22; -goto block_44; +x_27 = l_Lean_Meta_FVarSubst_apply(x_3, x_1); +if (lean_obj_tag(x_27) == 1) +{ +lean_object* x_46; +x_46 = lean_ctor_get(x_27, 0); +lean_inc(x_46); +lean_dec(x_27); +lean_ctor_set(x_23, 0, x_46); +x_28 = x_23; +goto block_45; } 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_free_object(x_22); -x_46 = l_Lean_MessageData_ofExpr(x_26); -x_47 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -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_RCases_rcasesCore___rarg___lambda__4___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_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_50, x_10, x_11, x_12, x_13, x_24); -x_27 = x_51; -goto block_44; +lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_free_object(x_23); +x_47 = l_Lean_MessageData_ofExpr(x_27); +x_48 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_49 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_47); +x_50 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___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_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_51, x_10, x_11, x_12, x_13, x_25); +x_28 = x_52; +goto block_45; } -block_44: +block_45: { -if (lean_obj_tag(x_27) == 0) +if (lean_obj_tag(x_28) == 0) { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_27, 0); -lean_inc(x_28); -x_29 = lean_ctor_get(x_27, 1); +lean_object* x_29; lean_object* x_30; lean_object* x_31; +x_29 = lean_ctor_get(x_28, 0); lean_inc(x_29); -lean_dec(x_27); +x_30 = lean_ctor_get(x_28, 1); +lean_inc(x_30); +lean_dec(x_28); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_30 = l_Lean_Elab_Tactic_RCases_subst_x27(x_4, x_28, x_3, x_10, x_11, x_12, x_13, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = l_Lean_Elab_Tactic_RCases_subst_x27(x_4, x_29, x_3, x_10, x_11, x_12, x_13, x_30); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_31 = lean_ctor_get(x_30, 0); -lean_inc(x_31); -x_32 = lean_ctor_get(x_30, 1); +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); -lean_dec(x_30); -x_33 = lean_ctor_get(x_31, 0); +x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_31, 1); -lean_inc(x_34); lean_dec(x_31); -x_35 = lean_apply_11(x_5, x_34, x_33, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_32); -return x_35; +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 = lean_apply_11(x_5, x_35, x_34, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_33); +return x_36; } else { -uint8_t x_36; +uint8_t x_37; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7243,29 +7244,29 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_36 = !lean_is_exclusive(x_30); -if (x_36 == 0) +x_37 = !lean_is_exclusive(x_31); +if (x_37 == 0) { -return x_30; +return x_31; } else { -lean_object* x_37; lean_object* x_38; lean_object* x_39; -x_37 = lean_ctor_get(x_30, 0); -x_38 = lean_ctor_get(x_30, 1); +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_inc(x_37); -lean_dec(x_30); -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_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 { -uint8_t x_40; +uint8_t x_41; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7277,96 +7278,96 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_40 = !lean_is_exclusive(x_27); -if (x_40 == 0) +x_41 = !lean_is_exclusive(x_28); +if (x_41 == 0) { -return x_27; +return x_28; } else { -lean_object* x_41; lean_object* x_42; lean_object* x_43; -x_41 = lean_ctor_get(x_27, 0); -x_42 = lean_ctor_get(x_27, 1); +lean_object* x_42; lean_object* x_43; lean_object* x_44; +x_42 = lean_ctor_get(x_28, 0); +x_43 = lean_ctor_get(x_28, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_27); -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_dec(x_28); +x_44 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_44, 0, x_42); +lean_ctor_set(x_44, 1, x_43); +return x_44; } } } } else { -lean_object* x_52; lean_object* x_53; lean_object* x_54; -x_52 = lean_ctor_get(x_22, 1); -lean_inc(x_52); -lean_dec(x_22); -x_53 = l_Lean_Meta_FVarSubst_apply(x_3, x_1); -if (lean_obj_tag(x_53) == 1) +lean_object* x_53; lean_object* x_54; lean_object* x_55; +x_53 = lean_ctor_get(x_23, 1); +lean_inc(x_53); +lean_dec(x_23); +x_54 = l_Lean_Meta_FVarSubst_apply(x_3, x_1); +if (lean_obj_tag(x_54) == 1) { -lean_object* x_72; lean_object* x_73; -x_72 = lean_ctor_get(x_53, 0); -lean_inc(x_72); -lean_dec(x_53); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_72); -lean_ctor_set(x_73, 1, x_52); -x_54 = x_73; -goto block_71; +lean_object* x_73; lean_object* x_74; +x_73 = lean_ctor_get(x_54, 0); +lean_inc(x_73); +lean_dec(x_54); +x_74 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_74, 0, x_73); +lean_ctor_set(x_74, 1, x_53); +x_55 = x_74; +goto block_72; } else { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_74 = l_Lean_MessageData_ofExpr(x_53); -x_75 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_76 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_74); -x_77 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; -x_78 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_78, 0, x_76); -lean_ctor_set(x_78, 1, x_77); -x_79 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_78, x_10, x_11, x_12, x_13, x_52); -x_54 = x_79; -goto block_71; +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 = l_Lean_MessageData_ofExpr(x_54); +x_76 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_77 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_75); +x_78 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_79 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_79, 0, x_77); +lean_ctor_set(x_79, 1, x_78); +x_80 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_79, x_10, x_11, x_12, x_13, x_53); +x_55 = x_80; +goto block_72; } -block_71: +block_72: { -if (lean_obj_tag(x_54) == 0) +if (lean_obj_tag(x_55) == 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_object* x_56; lean_object* x_57; lean_object* x_58; +x_56 = lean_ctor_get(x_55, 0); lean_inc(x_56); -lean_dec(x_54); +x_57 = lean_ctor_get(x_55, 1); +lean_inc(x_57); +lean_dec(x_55); lean_inc(x_13); lean_inc(x_12); lean_inc(x_11); lean_inc(x_10); -x_57 = l_Lean_Elab_Tactic_RCases_subst_x27(x_4, x_55, x_3, x_10, x_11, x_12, x_13, x_56); -if (lean_obj_tag(x_57) == 0) +x_58 = l_Lean_Elab_Tactic_RCases_subst_x27(x_4, x_56, x_3, x_10, x_11, x_12, x_13, x_57); +if (lean_obj_tag(x_58) == 0) { -lean_object* x_58; lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; -x_58 = lean_ctor_get(x_57, 0); -lean_inc(x_58); -x_59 = lean_ctor_get(x_57, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; lean_object* x_62; lean_object* x_63; +x_59 = lean_ctor_get(x_58, 0); lean_inc(x_59); -lean_dec(x_57); -x_60 = lean_ctor_get(x_58, 0); +x_60 = lean_ctor_get(x_58, 1); lean_inc(x_60); -x_61 = lean_ctor_get(x_58, 1); -lean_inc(x_61); lean_dec(x_58); -x_62 = lean_apply_11(x_5, x_61, x_60, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_59); -return x_62; +x_61 = lean_ctor_get(x_59, 0); +lean_inc(x_61); +x_62 = lean_ctor_get(x_59, 1); +lean_inc(x_62); +lean_dec(x_59); +x_63 = lean_apply_11(x_5, x_62, x_61, x_6, x_7, x_8, x_9, x_10, x_11, x_12, x_13, x_60); +return x_63; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7376,31 +7377,31 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); -x_63 = lean_ctor_get(x_57, 0); -lean_inc(x_63); -x_64 = lean_ctor_get(x_57, 1); +x_64 = lean_ctor_get(x_58, 0); lean_inc(x_64); -if (lean_is_exclusive(x_57)) { - lean_ctor_release(x_57, 0); - lean_ctor_release(x_57, 1); - x_65 = x_57; +x_65 = lean_ctor_get(x_58, 1); +lean_inc(x_65); +if (lean_is_exclusive(x_58)) { + lean_ctor_release(x_58, 0); + lean_ctor_release(x_58, 1); + x_66 = x_58; } else { - lean_dec_ref(x_57); - x_65 = lean_box(0); + lean_dec_ref(x_58); + x_66 = lean_box(0); } -if (lean_is_scalar(x_65)) { - x_66 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_66)) { + x_67 = lean_alloc_ctor(1, 2, 0); } else { - x_66 = x_65; + x_67 = x_66; } -lean_ctor_set(x_66, 0, x_63); -lean_ctor_set(x_66, 1, x_64); -return x_66; +lean_ctor_set(x_67, 0, x_64); +lean_ctor_set(x_67, 1, x_65); +return x_67; } } else { -lean_object* x_67; lean_object* x_68; lean_object* x_69; lean_object* x_70; +lean_object* x_68; lean_object* x_69; lean_object* x_70; lean_object* x_71; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7412,33 +7413,33 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_67 = lean_ctor_get(x_54, 0); -lean_inc(x_67); -x_68 = lean_ctor_get(x_54, 1); +x_68 = lean_ctor_get(x_55, 0); lean_inc(x_68); -if (lean_is_exclusive(x_54)) { - lean_ctor_release(x_54, 0); - lean_ctor_release(x_54, 1); - x_69 = x_54; +x_69 = lean_ctor_get(x_55, 1); +lean_inc(x_69); +if (lean_is_exclusive(x_55)) { + lean_ctor_release(x_55, 0); + lean_ctor_release(x_55, 1); + x_70 = x_55; } else { - lean_dec_ref(x_54); - x_69 = lean_box(0); + lean_dec_ref(x_55); + x_70 = lean_box(0); } -if (lean_is_scalar(x_69)) { - x_70 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_70)) { + x_71 = lean_alloc_ctor(1, 2, 0); } else { - x_70 = x_69; + x_71 = x_70; } -lean_ctor_set(x_70, 0, x_67); -lean_ctor_set(x_70, 1, x_68); -return x_70; +lean_ctor_set(x_71, 0, x_68); +lean_ctor_set(x_71, 1, x_69); +return x_71; } } } } else { -uint8_t x_80; +uint8_t x_81; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7451,29 +7452,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_80 = !lean_is_exclusive(x_22); -if (x_80 == 0) +x_81 = !lean_is_exclusive(x_23); +if (x_81 == 0) { -return x_22; +return x_23; } else { -lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_81 = lean_ctor_get(x_22, 0); -x_82 = lean_ctor_get(x_22, 1); +lean_object* x_82; lean_object* x_83; lean_object* x_84; +x_82 = lean_ctor_get(x_23, 0); +x_83 = lean_ctor_get(x_23, 1); +lean_inc(x_83); lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_22); -x_83 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_83, 0, x_81); -lean_ctor_set(x_83, 1, x_82); -return x_83; +lean_dec(x_23); +x_84 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_84, 0, x_82); +lean_ctor_set(x_84, 1, x_83); +return x_84; } } } else { -uint8_t x_84; +uint8_t x_85; lean_dec(x_13); lean_dec(x_12); lean_dec(x_11); @@ -7487,23 +7488,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_84 = !lean_is_exclusive(x_16); -if (x_84 == 0) +x_85 = !lean_is_exclusive(x_17); +if (x_85 == 0) { -return x_16; +return x_17; } 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_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_17, 0); +x_87 = lean_ctor_get(x_17, 1); +lean_inc(x_87); 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_dec(x_17); +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; } } } @@ -8806,148 +8807,149 @@ lean_dec(x_24); x_25 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); if (lean_obj_tag(x_25) == 1) { -lean_object* x_227; -x_227 = lean_ctor_get(x_25, 0); -lean_inc(x_227); -lean_ctor_set(x_21, 0, x_227); +lean_object* x_228; +x_228 = lean_ctor_get(x_25, 0); +lean_inc(x_228); +lean_ctor_set(x_21, 0, x_228); x_26 = x_21; -goto block_226; +goto block_227; } else { -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_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_free_object(x_21); lean_inc(x_25); -x_228 = l_Lean_MessageData_ofExpr(x_25); -x_229 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_230 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_228); -x_231 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; -x_232 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_232, 0, x_230); -lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_232, x_14, x_15, x_16, x_17, x_23); -x_26 = x_233; -goto block_226; -} -block_226: +x_229 = l_Lean_MessageData_ofExpr(x_25); +x_230 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_231 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_229); +x_232 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_233 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_233, x_14, x_15, x_16, x_17, x_23); +x_26 = x_234; +goto block_227; +} +block_227: { if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_27; uint8_t x_28; lean_object* x_29; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); +x_28 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_28 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -if (lean_obj_tag(x_28) == 0) +x_29 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_28, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_25); -x_30 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_33 = l_Lean_Meta_whnfD(x_31, x_14, x_15, x_16, x_17, x_32); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_Meta_whnfD(x_32, x_14, x_15, x_16, x_17, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Expr_getAppFn(x_34); -if (lean_obj_tag(x_36) == 4) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Expr_getAppFn(x_35); +if (lean_obj_tag(x_37) == 4) { -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_ctor_get(x_36, 0); -lean_inc(x_110); -lean_dec(x_36); -x_111 = lean_st_ref_get(x_17, x_35); -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_111 = lean_ctor_get(x_37, 0); +lean_inc(x_111); +lean_dec(x_37); +x_112 = lean_st_ref_get(x_17, x_36); +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -lean_dec(x_111); -x_114 = lean_ctor_get(x_112, 0); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_110); -if (lean_obj_tag(x_115) == 0) +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_environment_find(x_115, x_111); +if (lean_obj_tag(x_116) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_object* x_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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_116 = l_Lean_MessageData_ofExpr(x_25); -x_117 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_116); -x_119 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_MessageData_ofExpr(x_34); -x_122 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_124 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); +x_117 = l_Lean_MessageData_ofExpr(x_25); +x_118 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_119 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_121 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Lean_MessageData_ofExpr(x_35); +x_123 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_125 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); lean_inc(x_12); -x_125 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_124, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_125; -goto block_109; +x_126 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_125, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_126; +goto block_110; } else { -lean_object* x_126; -x_126 = lean_ctor_get(x_115, 0); -lean_inc(x_126); -lean_dec(x_115); -switch (lean_obj_tag(x_126)) { +lean_object* x_127; +x_127 = lean_ctor_get(x_116, 0); +lean_inc(x_127); +lean_dec(x_116); +switch (lean_obj_tag(x_127)) { case 4: { -lean_object* x_127; lean_object* x_128; uint8_t x_144; lean_object* x_145; +lean_object* x_128; lean_object* x_129; uint8_t x_145; lean_object* x_146; lean_dec(x_1); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -lean_dec(x_126); -x_144 = lean_ctor_get_uint8(x_127, sizeof(void*)*1); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); lean_dec(x_127); -x_145 = lean_box(x_144); -if (lean_obj_tag(x_145) == 0) +x_145 = lean_ctor_get_uint8(x_128, sizeof(void*)*1); +lean_dec(x_128); +x_146 = lean_box(x_145); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_34); -x_146 = lean_box(0); +lean_object* x_147; lean_object* x_148; +lean_dec(x_35); +x_147 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -8955,357 +8957,357 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_147 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_146, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_147; -goto block_109; +x_148 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_148; +goto block_110; } else { -lean_object* x_148; -lean_dec(x_145); +lean_object* x_149; +lean_dec(x_146); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_148 = lean_box(0); -x_128 = x_148; -goto block_143; +x_149 = lean_box(0); +x_129 = x_149; +goto block_144; } -block_143: +block_144: { -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; uint8_t x_139; -lean_dec(x_128); -x_129 = l_Lean_MessageData_ofExpr(x_25); -x_130 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_131 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_133 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_MessageData_ofExpr(x_34); -x_135 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_137 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); +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; uint8_t x_140; +lean_dec(x_129); +x_130 = l_Lean_MessageData_ofExpr(x_25); +x_131 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_132 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_130); +x_133 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_134 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_MessageData_ofExpr(x_35); +x_136 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_138 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); lean_inc(x_12); -x_138 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_137, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_139 = !lean_is_exclusive(x_138); -if (x_139 == 0) +x_139 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_138, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) { -x_37 = x_138; -goto block_109; +x_38 = x_139; +goto block_110; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_138, 0); -x_141 = lean_ctor_get(x_138, 1); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_139, 0); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_138); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_37 = x_142; -goto block_109; +lean_dec(x_139); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_38 = x_143; +goto block_110; } } } case 5: { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_34); +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_35); lean_dec(x_11); lean_dec(x_3); -x_149 = lean_ctor_get(x_126, 0); -lean_inc(x_149); -lean_dec(x_126); -x_150 = lean_ctor_get(x_149, 1); +x_150 = lean_ctor_get(x_127, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 4); +lean_dec(x_127); +x_151 = lean_ctor_get(x_150, 1); lean_inc(x_151); -lean_dec(x_149); -x_152 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_153 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +x_152 = lean_ctor_get(x_150, 4); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_154 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_154 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_150, x_153, x_151, x_152, x_14, x_15, x_16, x_17, x_113); -lean_dec(x_150); -if (lean_obj_tag(x_154) == 0) +x_155 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_151, x_154, x_152, x_153, x_14, x_15, x_16, x_17, x_114); +lean_dec(x_151); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); -lean_dec(x_154); -x_157 = !lean_is_exclusive(x_155); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; -x_158 = lean_ctor_get(x_155, 0); -x_159 = lean_ctor_get(x_155, 1); -x_160 = l_Lean_Expr_fvarId_x21(x_25); -x_161 = 1; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_158 = !lean_is_exclusive(x_156); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; +x_159 = lean_ctor_get(x_156, 0); +x_160 = lean_ctor_get(x_156, 1); +x_161 = l_Lean_Expr_fvarId_x21(x_25); +x_162 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_162 = l_Lean_Meta_Cases_cases(x_10, x_160, x_158, x_161, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_162) == 0) +x_163 = l_Lean_Meta_Cases_cases(x_10, x_161, x_159, x_162, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_163) == 0) { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_162); -if (x_163 == 0) +uint8_t x_164; +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) { -lean_object* x_164; -x_164 = lean_ctor_get(x_162, 0); -lean_ctor_set(x_155, 1, x_164); -lean_ctor_set(x_155, 0, x_159); -lean_ctor_set(x_162, 0, x_155); -x_37 = x_162; -goto block_109; +lean_object* x_165; +x_165 = lean_ctor_get(x_163, 0); +lean_ctor_set(x_156, 1, x_165); +lean_ctor_set(x_156, 0, x_160); +lean_ctor_set(x_163, 0, x_156); +x_38 = x_163; +goto block_110; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_162, 0); -x_166 = lean_ctor_get(x_162, 1); +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_163, 0); +x_167 = lean_ctor_get(x_163, 1); +lean_inc(x_167); lean_inc(x_166); -lean_inc(x_165); -lean_dec(x_162); -lean_ctor_set(x_155, 1, x_165); -lean_ctor_set(x_155, 0, x_159); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_155); -lean_ctor_set(x_167, 1, x_166); -x_37 = x_167; -goto block_109; +lean_dec(x_163); +lean_ctor_set(x_156, 1, x_166); +lean_ctor_set(x_156, 0, x_160); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_156); +lean_ctor_set(x_168, 1, x_167); +x_38 = x_168; +goto block_110; } } else { -uint8_t x_168; -lean_free_object(x_155); -lean_dec(x_159); -x_168 = !lean_is_exclusive(x_162); -if (x_168 == 0) +uint8_t x_169; +lean_free_object(x_156); +lean_dec(x_160); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -x_37 = x_162; -goto block_109; +x_38 = x_163; +goto block_110; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_162, 0); -x_170 = lean_ctor_get(x_162, 1); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_163, 0); +x_171 = lean_ctor_get(x_163, 1); +lean_inc(x_171); lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_162); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_37 = x_171; -goto block_109; +lean_dec(x_163); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_38 = x_172; +goto block_110; } } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; -x_172 = lean_ctor_get(x_155, 0); -x_173 = lean_ctor_get(x_155, 1); +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; +x_173 = lean_ctor_get(x_156, 0); +x_174 = lean_ctor_get(x_156, 1); +lean_inc(x_174); lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_155); -x_174 = l_Lean_Expr_fvarId_x21(x_25); -x_175 = 1; +lean_dec(x_156); +x_175 = l_Lean_Expr_fvarId_x21(x_25); +x_176 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_176 = l_Lean_Meta_Cases_cases(x_10, x_174, x_172, x_175, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_176) == 0) +x_177 = l_Lean_Meta_Cases_cases(x_10, x_175, x_173, x_176, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_177) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_179 = x_176; +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_180 = x_177; } else { - lean_dec_ref(x_176); - x_179 = lean_box(0); + lean_dec_ref(x_177); + x_180 = lean_box(0); } -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_173); -lean_ctor_set(x_180, 1, x_177); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(0, 2, 0); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_174); +lean_ctor_set(x_181, 1, x_178); +if (lean_is_scalar(x_180)) { + x_182 = lean_alloc_ctor(0, 2, 0); } else { - x_181 = x_179; + x_182 = x_180; } -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_178); -x_37 = x_181; -goto block_109; +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +x_38 = x_182; +goto block_110; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_173); -x_182 = lean_ctor_get(x_176, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_176, 1); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_174); +x_183 = lean_ctor_get(x_177, 0); lean_inc(x_183); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_184 = x_176; +x_184 = lean_ctor_get(x_177, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_185 = x_177; } else { - lean_dec_ref(x_176); - x_184 = lean_box(0); + lean_dec_ref(x_177); + x_185 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_186 = x_185; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -x_37 = x_185; -goto block_109; +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_184); +x_38 = x_186; +goto block_110; } } } else { -uint8_t x_186; +uint8_t x_187; lean_dec(x_25); lean_dec(x_10); -x_186 = !lean_is_exclusive(x_154); -if (x_186 == 0) +x_187 = !lean_is_exclusive(x_155); +if (x_187 == 0) { -x_37 = x_154; -goto block_109; +x_38 = x_155; +goto block_110; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_154, 0); -x_188 = lean_ctor_get(x_154, 1); +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_155, 0); +x_189 = lean_ctor_get(x_155, 1); +lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_154); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -x_37 = x_189; -goto block_109; +lean_dec(x_155); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_38 = x_190; +goto block_110; } } } default: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_126); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_127); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_190 = l_Lean_MessageData_ofExpr(x_25); -x_191 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_192 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_190); -x_193 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_MessageData_ofExpr(x_34); -x_196 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -x_197 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_198 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); +x_191 = l_Lean_MessageData_ofExpr(x_25); +x_192 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_MessageData_ofExpr(x_35); +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 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_199 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); lean_inc(x_12); -x_199 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_199; -goto block_109; +x_200 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_199, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_200; +goto block_110; } } } } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_36); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_37); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_200 = l_Lean_MessageData_ofExpr(x_25); -x_201 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_202 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_200); -x_203 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_204 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -x_205 = l_Lean_MessageData_ofExpr(x_34); -x_206 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_208 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); +x_201 = l_Lean_MessageData_ofExpr(x_25); +x_202 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_203 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_205 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_MessageData_ofExpr(x_35); +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +x_208 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_209 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); lean_inc(x_12); -x_209 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_208, x_12, x_13, x_14, x_15, x_16, x_17, x_35); -x_37 = x_209; -goto block_109; +x_210 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_209, x_12, x_13, x_14, x_15, x_16, x_17, x_36); +x_38 = x_210; +goto block_110; } -block_109: +block_110: { -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_38) == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_38, 0); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_37, 1); -x_42 = lean_ctor_get(x_39, 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_38, 1); +x_43 = lean_ctor_get(x_40, 1); lean_inc(x_6); -lean_ctor_set(x_39, 1, x_6); -x_43 = lean_array_get_size(x_42); -x_44 = lean_unsigned_to_nat(0u); -x_45 = lean_nat_dec_lt(x_44, x_43); -if (x_45 == 0) +lean_ctor_set(x_40, 1, x_6); +x_44 = lean_array_get_size(x_43); +x_45 = lean_unsigned_to_nat(0u); +x_46 = lean_nat_dec_lt(x_45, x_44); +if (x_46 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9315,18 +9317,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_43, x_43); -if (x_46 == 0) +uint8_t x_47; +x_47 = lean_nat_dec_le(x_44, x_44); +if (x_47 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9336,70 +9338,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_47; size_t x_48; lean_object* x_49; -lean_free_object(x_37); +size_t x_48; size_t x_49; lean_object* x_50; +lean_free_object(x_38); lean_dec(x_6); -x_47 = 0; -x_48 = lean_usize_of_nat(x_43); +x_48 = 0; +x_49 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_43, x_48, x_49, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_42); lean_dec(x_43); -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_42, x_47, x_48, x_39, x_12, x_13, x_14, x_15, x_16, x_17, x_41); -lean_dec(x_42); -if (lean_obj_tag(x_49) == 0) +if (lean_obj_tag(x_50) == 0) { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_52); -return x_49; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_53); +return x_50; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_49, 0); -x_54 = lean_ctor_get(x_49, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_49); -x_55 = lean_ctor_get(x_53, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +x_55 = lean_ctor_get(x_50, 1); lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_inc(x_54); +lean_dec(x_50); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_58; +x_58 = !lean_is_exclusive(x_50); +if (x_58 == 0) { -return x_49; +return x_50; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_50, 0); +x_60 = lean_ctor_get(x_50, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_50); +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; } } } @@ -9407,25 +9409,25 @@ return x_60; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_61 = lean_ctor_get(x_37, 1); -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_62 = lean_ctor_get(x_38, 1); +x_63 = lean_ctor_get(x_40, 0); +x_64 = lean_ctor_get(x_40, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); +lean_dec(x_40); lean_inc(x_6); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_6); -x_65 = lean_array_get_size(x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_6); +x_66 = lean_array_get_size(x_64); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +if (x_68 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9435,18 +9437,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_68; -x_68 = lean_nat_dec_le(x_65, x_65); -if (x_68 == 0) +uint8_t x_69; +x_69 = lean_nat_dec_le(x_66, x_66); +if (x_69 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9456,69 +9458,69 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_69; size_t x_70; lean_object* x_71; -lean_free_object(x_37); +size_t x_70; size_t x_71; lean_object* x_72; +lean_free_object(x_38); lean_dec(x_6); -x_69 = 0; -x_70 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_71 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_63, x_69, x_70, x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -lean_dec(x_63); -if (lean_obj_tag(x_71) == 0) +x_70 = 0; +x_71 = lean_usize_of_nat(x_66); +lean_dec(x_66); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_64, x_70, x_71, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_64); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); 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_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_72); + x_75 = lean_box(0); } -x_75 = lean_ctor_get(x_72, 1); -lean_inc(x_75); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_76 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_76 = x_74; + x_77 = x_75; } -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_73); -return x_76; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_71, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_71, 1); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_72, 0); lean_inc(x_78); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_79 = x_71; +x_79 = lean_ctor_get(x_72, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_80 = x_72; } else { - lean_dec_ref(x_71); - x_79 = lean_box(0); + lean_dec_ref(x_72); + x_80 = lean_box(0); } -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_80 = x_79; + x_81 = x_80; } -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_78); -return x_80; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } @@ -9526,41 +9528,41 @@ 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; uint8_t x_89; -x_81 = lean_ctor_get(x_37, 0); -x_82 = lean_ctor_get(x_37, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_37); -x_83 = lean_ctor_get(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; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_82 = lean_ctor_get(x_38, 0); +x_83 = lean_ctor_get(x_38, 1); lean_inc(x_83); -x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_38); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_85 = x_81; +x_85 = lean_ctor_get(x_82, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_81); - x_85 = lean_box(0); + lean_dec_ref(x_82); + x_86 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_86 = x_85; + x_87 = x_86; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_6); -x_87 = lean_array_get_size(x_84); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_nat_dec_lt(x_88, x_87); -if (x_89 == 0) +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_6); +x_88 = lean_array_get_size(x_85); +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_nat_dec_lt(x_89, x_88); +if (x_90 == 0) { -lean_object* x_90; +lean_object* x_91; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9570,21 +9572,21 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_82); -return x_90; +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_6); +lean_ctor_set(x_91, 1, x_83); +return x_91; } else { -uint8_t x_91; -x_91 = lean_nat_dec_le(x_87, x_87); -if (x_91 == 0) +uint8_t x_92; +x_92 = lean_nat_dec_le(x_88, x_88); +if (x_92 == 0) { -lean_object* x_92; +lean_object* x_93; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9594,70 +9596,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_6); -lean_ctor_set(x_92, 1, x_82); -return x_92; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_6); +lean_ctor_set(x_93, 1, x_83); +return x_93; } else { -size_t x_93; size_t x_94; lean_object* x_95; +size_t x_94; size_t x_95; lean_object* x_96; lean_dec(x_6); -x_93 = 0; -x_94 = lean_usize_of_nat(x_87); -lean_dec(x_87); -x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_84, x_93, x_94, x_86, x_12, x_13, x_14, x_15, x_16, x_17, x_82); -lean_dec(x_84); -if (lean_obj_tag(x_95) == 0) +x_94 = 0; +x_95 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_96 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_85, x_94, x_95, x_87, x_12, x_13, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_85); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_ctor_get(x_96, 0); lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_99 = x_96; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_96); + x_99 = lean_box(0); } -x_99 = lean_ctor_get(x_96, 1); -lean_inc(x_99); -lean_dec(x_96); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_101 = x_99; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_101 = lean_ctor_get(x_95, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_95, 1); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_96, 0); lean_inc(x_102); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_103 = x_95; +x_103 = lean_ctor_get(x_96, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_104 = x_96; } else { - lean_dec_ref(x_95); - x_103 = lean_box(0); + lean_dec_ref(x_96); + x_104 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_105 = x_104; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } } @@ -9665,7 +9667,7 @@ return x_104; } else { -uint8_t x_105; +uint8_t x_106; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -9676,30 +9678,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_37); -if (x_105 == 0) +x_106 = !lean_is_exclusive(x_38); +if (x_106 == 0) { -return x_37; +return x_38; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_37, 0); -x_107 = lean_ctor_get(x_37, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_38, 0); +x_108 = lean_ctor_get(x_38, 1); +lean_inc(x_108); lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_37); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_dec(x_38); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; } } } } else { -uint8_t x_210; +uint8_t x_211; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -9716,29 +9718,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_210 = !lean_is_exclusive(x_33); -if (x_210 == 0) +x_211 = !lean_is_exclusive(x_34); +if (x_211 == 0) { -return x_33; +return x_34; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_33, 0); -x_212 = lean_ctor_get(x_33, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_34, 0); +x_213 = lean_ctor_get(x_34, 1); +lean_inc(x_213); lean_inc(x_212); -lean_inc(x_211); -lean_dec(x_33); -x_213 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -return x_213; +lean_dec(x_34); +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +return x_214; } } } else { -uint8_t x_214; +uint8_t x_215; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -9755,29 +9757,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_214 = !lean_is_exclusive(x_30); -if (x_214 == 0) +x_215 = !lean_is_exclusive(x_31); +if (x_215 == 0) { -return x_30; +return x_31; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_215 = lean_ctor_get(x_30, 0); -x_216 = lean_ctor_get(x_30, 1); +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_ctor_get(x_31, 0); +x_217 = lean_ctor_get(x_31, 1); +lean_inc(x_217); lean_inc(x_216); -lean_inc(x_215); -lean_dec(x_30); -x_217 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -return x_217; +lean_dec(x_31); +x_218 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +return x_218; } } } else { -uint8_t x_218; +uint8_t x_219; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -9794,29 +9796,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_218 = !lean_is_exclusive(x_28); -if (x_218 == 0) +x_219 = !lean_is_exclusive(x_29); +if (x_219 == 0) { -return x_28; +return x_29; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_219 = lean_ctor_get(x_28, 0); -x_220 = lean_ctor_get(x_28, 1); +lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_220 = lean_ctor_get(x_29, 0); +x_221 = lean_ctor_get(x_29, 1); +lean_inc(x_221); lean_inc(x_220); -lean_inc(x_219); -lean_dec(x_28); -x_221 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -return x_221; +lean_dec(x_29); +x_222 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } } else { -uint8_t x_222; +uint8_t x_223; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -9833,179 +9835,180 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_222 = !lean_is_exclusive(x_26); -if (x_222 == 0) +x_223 = !lean_is_exclusive(x_26); +if (x_223 == 0) { return x_26; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_26, 0); -x_224 = lean_ctor_get(x_26, 1); +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_26, 0); +x_225 = lean_ctor_get(x_26, 1); +lean_inc(x_225); lean_inc(x_224); -lean_inc(x_223); lean_dec(x_26); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +x_226 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_226, 0, x_224); +lean_ctor_set(x_226, 1, x_225); +return x_226; } } } } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_234 = lean_ctor_get(x_21, 1); -lean_inc(x_234); -lean_dec(x_21); -x_235 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); -if (lean_obj_tag(x_235) == 1) +lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_235 = lean_ctor_get(x_21, 1); +lean_inc(x_235); +lean_dec(x_21); +x_236 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); +if (lean_obj_tag(x_236) == 1) { -lean_object* x_381; lean_object* x_382; -x_381 = lean_ctor_get(x_235, 0); -lean_inc(x_381); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_381); -lean_ctor_set(x_382, 1, x_234); -x_236 = x_382; -goto block_380; +lean_object* x_383; lean_object* x_384; +x_383 = lean_ctor_get(x_236, 0); +lean_inc(x_383); +x_384 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_235); +x_237 = x_384; +goto block_382; } else { -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_inc(x_235); -x_383 = l_Lean_MessageData_ofExpr(x_235); -x_384 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_385 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_383); -x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +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_inc(x_236); +x_385 = l_Lean_MessageData_ofExpr(x_236); +x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_387 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_387, 0, x_385); -lean_ctor_set(x_387, 1, x_386); -x_388 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_387, x_14, x_15, x_16, x_17, x_234); -x_236 = x_388; -goto block_380; -} -block_380: -{ -if (lean_obj_tag(x_236) == 0) -{ -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_236, 1); -lean_inc(x_237); -lean_dec(x_236); +lean_ctor_set(x_387, 0, x_386); +lean_ctor_set(x_387, 1, x_385); +x_388 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_389 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_389, 0, x_387); +lean_ctor_set(x_389, 1, x_388); +x_390 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_389, x_14, x_15, x_16, x_17, x_235); +x_237 = x_390; +goto block_382; +} +block_382: +{ +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; uint8_t x_239; lean_object* x_240; +x_238 = lean_ctor_get(x_237, 1); +lean_inc(x_238); +lean_dec(x_237); +x_239 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_238 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_237); -if (lean_obj_tag(x_238) == 0) -{ -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_238, 1); -lean_inc(x_239); -lean_dec(x_238); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_235); -x_240 = lean_infer_type(x_235, x_14, x_15, x_16, x_17, x_239); +x_240 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_239, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_238); if (lean_obj_tag(x_240) == 0) { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_240, 0); +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_240, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); lean_dec(x_240); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_243 = l_Lean_Meta_whnfD(x_241, x_14, x_15, x_16, x_17, x_242); -if (lean_obj_tag(x_243) == 0) +lean_inc(x_236); +x_242 = lean_infer_type(x_236, x_14, x_15, x_16, x_17, x_241); +if (lean_obj_tag(x_242) == 0) { -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_244 = lean_ctor_get(x_243, 0); +lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); -x_245 = lean_ctor_get(x_243, 1); -lean_inc(x_245); -lean_dec(x_243); -x_246 = l_Lean_Expr_getAppFn(x_244); -if (lean_obj_tag(x_246) == 4) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_278 = lean_ctor_get(x_246, 0); -lean_inc(x_278); -lean_dec(x_246); -x_279 = lean_st_ref_get(x_17, x_245); -x_280 = lean_ctor_get(x_279, 0); +lean_dec(x_242); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_245 = l_Lean_Meta_whnfD(x_243, x_14, x_15, x_16, x_17, x_244); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec(x_245); +x_248 = l_Lean_Expr_getAppFn(x_246); +if (lean_obj_tag(x_248) == 4) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_280 = lean_ctor_get(x_248, 0); lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_282 = lean_ctor_get(x_280, 0); +lean_dec(x_248); +x_281 = lean_st_ref_get(x_17, x_247); +x_282 = lean_ctor_get(x_281, 0); lean_inc(x_282); -lean_dec(x_280); -x_283 = lean_environment_find(x_282, x_278); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_ctor_get(x_282, 0); +lean_inc(x_284); +lean_dec(x_282); +x_285 = lean_environment_find(x_284, x_280); +if (lean_obj_tag(x_285) == 0) +{ +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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_284 = l_Lean_MessageData_ofExpr(x_235); -x_285 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_286 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_284); -x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_286 = l_Lean_MessageData_ofExpr(x_236); +x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_288 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_288, 0, x_286); -lean_ctor_set(x_288, 1, x_287); -x_289 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_286); +x_289 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_290 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_290, 0, x_288); lean_ctor_set(x_290, 1, x_289); -x_291 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_291 = l_Lean_MessageData_ofExpr(x_246); x_292 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_292, 0, x_290); lean_ctor_set(x_292, 1, x_291); +x_293 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_294 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); lean_inc(x_12); -x_293 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_293; -goto block_277; +x_295 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_294, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_295; +goto block_279; } else { -lean_object* x_294; -x_294 = lean_ctor_get(x_283, 0); -lean_inc(x_294); -lean_dec(x_283); -switch (lean_obj_tag(x_294)) { +lean_object* x_296; +x_296 = lean_ctor_get(x_285, 0); +lean_inc(x_296); +lean_dec(x_285); +switch (lean_obj_tag(x_296)) { case 4: { -lean_object* x_295; lean_object* x_296; uint8_t x_312; lean_object* x_313; +lean_object* x_297; lean_object* x_298; uint8_t x_314; lean_object* x_315; lean_dec(x_1); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -lean_dec(x_294); -x_312 = lean_ctor_get_uint8(x_295, sizeof(void*)*1); -lean_dec(x_295); -x_313 = lean_box(x_312); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; lean_object* x_315; -lean_dec(x_244); -x_314 = lean_box(0); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +lean_dec(x_296); +x_314 = lean_ctor_get_uint8(x_297, sizeof(void*)*1); +lean_dec(x_297); +x_315 = lean_box(x_314); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; +lean_dec(x_246); +x_316 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -10013,316 +10016,316 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_315 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_235, x_10, x_3, x_11, x_4, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_315; -goto block_277; +x_317 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_236, x_10, x_3, x_11, x_4, x_316, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_317; +goto block_279; } else { -lean_object* x_316; -lean_dec(x_313); +lean_object* x_318; +lean_dec(x_315); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_316 = lean_box(0); -x_296 = x_316; -goto block_311; +x_318 = lean_box(0); +x_298 = x_318; +goto block_313; } -block_311: +block_313: { -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_296); -x_297 = l_Lean_MessageData_ofExpr(x_235); -x_298 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_299 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +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_dec(x_298); +x_299 = l_Lean_MessageData_ofExpr(x_236); +x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_301 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_299); +x_302 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_303 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_303, 0, x_301); lean_ctor_set(x_303, 1, x_302); -x_304 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_304 = l_Lean_MessageData_ofExpr(x_246); x_305 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_305, 0, x_303); lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_307 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_307, 0, x_305); +lean_ctor_set(x_307, 1, x_306); lean_inc(x_12); -x_306 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_305, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_307 = lean_ctor_get(x_306, 0); -lean_inc(x_307); -x_308 = lean_ctor_get(x_306, 1); -lean_inc(x_308); -if (lean_is_exclusive(x_306)) { - lean_ctor_release(x_306, 0); - lean_ctor_release(x_306, 1); - x_309 = x_306; +x_308 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_307, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_311 = x_308; } else { - lean_dec_ref(x_306); - x_309 = lean_box(0); + lean_dec_ref(x_308); + x_311 = lean_box(0); } -if (lean_is_scalar(x_309)) { - x_310 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_311)) { + x_312 = lean_alloc_ctor(1, 2, 0); } else { - x_310 = x_309; + x_312 = x_311; } -lean_ctor_set(x_310, 0, x_307); -lean_ctor_set(x_310, 1, x_308); -x_247 = x_310; -goto block_277; +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_310); +x_249 = x_312; +goto block_279; } } case 5: { -lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_244); +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_246); lean_dec(x_11); lean_dec(x_3); -x_317 = lean_ctor_get(x_294, 0); -lean_inc(x_317); -lean_dec(x_294); -x_318 = lean_ctor_get(x_317, 1); -lean_inc(x_318); -x_319 = lean_ctor_get(x_317, 4); +x_319 = lean_ctor_get(x_296, 0); lean_inc(x_319); -lean_dec(x_317); -x_320 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_321 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +lean_dec(x_296); +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +x_321 = lean_ctor_get(x_319, 4); +lean_inc(x_321); +lean_dec(x_319); +x_322 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_323 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_322 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_318, x_321, x_319, x_320, x_14, x_15, x_16, x_17, x_281); -lean_dec(x_318); -if (lean_obj_tag(x_322) == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; lean_object* x_330; -x_323 = lean_ctor_get(x_322, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_322, 1); -lean_inc(x_324); -lean_dec(x_322); -x_325 = lean_ctor_get(x_323, 0); +x_324 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_320, x_323, x_321, x_322, x_14, x_15, x_16, x_17, x_283); +lean_dec(x_320); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; lean_object* x_332; +x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); -x_326 = lean_ctor_get(x_323, 1); +x_326 = lean_ctor_get(x_324, 1); lean_inc(x_326); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_327 = x_323; -} else { - lean_dec_ref(x_323); - x_327 = lean_box(0); -} -x_328 = l_Lean_Expr_fvarId_x21(x_235); -x_329 = 1; +lean_dec(x_324); +x_327 = lean_ctor_get(x_325, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_329 = x_325; +} else { + lean_dec_ref(x_325); + x_329 = lean_box(0); +} +x_330 = l_Lean_Expr_fvarId_x21(x_236); +x_331 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_330 = l_Lean_Meta_Cases_cases(x_10, x_328, x_325, x_329, x_14, x_15, x_16, x_17, x_324); -if (lean_obj_tag(x_330) == 0) +x_332 = l_Lean_Meta_Cases_cases(x_10, x_330, x_327, x_331, x_14, x_15, x_16, x_17, x_326); +if (lean_obj_tag(x_332) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_330, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_333 = x_330; +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_332, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_335 = x_332; } else { - lean_dec_ref(x_330); - x_333 = lean_box(0); + lean_dec_ref(x_332); + x_335 = lean_box(0); } -if (lean_is_scalar(x_327)) { - x_334 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_329)) { + x_336 = lean_alloc_ctor(0, 2, 0); } else { - x_334 = x_327; + x_336 = x_329; } -lean_ctor_set(x_334, 0, x_326); -lean_ctor_set(x_334, 1, x_331); -if (lean_is_scalar(x_333)) { - x_335 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_328); +lean_ctor_set(x_336, 1, x_333); +if (lean_is_scalar(x_335)) { + x_337 = lean_alloc_ctor(0, 2, 0); } else { - x_335 = x_333; + x_337 = x_335; } -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_332); -x_247 = x_335; -goto block_277; +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_334); +x_249 = x_337; +goto block_279; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_327); -lean_dec(x_326); -x_336 = lean_ctor_get(x_330, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_330, 1); -lean_inc(x_337); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_338 = x_330; +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_329); +lean_dec(x_328); +x_338 = lean_ctor_get(x_332, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_332, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_340 = x_332; } else { - lean_dec_ref(x_330); - x_338 = lean_box(0); + lean_dec_ref(x_332); + x_340 = lean_box(0); } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); } else { - x_339 = x_338; + x_341 = x_340; } -lean_ctor_set(x_339, 0, x_336); -lean_ctor_set(x_339, 1, x_337); -x_247 = x_339; -goto block_277; +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +x_249 = x_341; +goto block_279; } } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_235); +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_236); lean_dec(x_10); -x_340 = lean_ctor_get(x_322, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_322, 1); -lean_inc(x_341); -if (lean_is_exclusive(x_322)) { - lean_ctor_release(x_322, 0); - lean_ctor_release(x_322, 1); - x_342 = x_322; +x_342 = lean_ctor_get(x_324, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_324, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_344 = x_324; } else { - lean_dec_ref(x_322); - x_342 = lean_box(0); + lean_dec_ref(x_324); + x_344 = lean_box(0); } -if (lean_is_scalar(x_342)) { - x_343 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); } else { - x_343 = x_342; + x_345 = x_344; } -lean_ctor_set(x_343, 0, x_340); -lean_ctor_set(x_343, 1, x_341); -x_247 = x_343; -goto block_277; +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +x_249 = x_345; +goto block_279; } } default: { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -lean_dec(x_294); +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_296); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_344 = l_Lean_MessageData_ofExpr(x_235); -x_345 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_346 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_344); -x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_346 = l_Lean_MessageData_ofExpr(x_236); +x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_348 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_348, 0, x_346); -lean_ctor_set(x_348, 1, x_347); -x_349 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_348, 0, x_347); +lean_ctor_set(x_348, 1, x_346); +x_349 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_350 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_350, 0, x_348); lean_ctor_set(x_350, 1, x_349); -x_351 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_351 = l_Lean_MessageData_ofExpr(x_246); x_352 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_352, 0, x_350); lean_ctor_set(x_352, 1, x_351); +x_353 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_354 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); lean_inc(x_12); -x_353 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_352, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_353; -goto block_277; +x_355 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_354, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_355; +goto block_279; } } } } 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; -lean_dec(x_246); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; +lean_dec(x_248); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_354 = l_Lean_MessageData_ofExpr(x_235); -x_355 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_356 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_354); -x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_356 = l_Lean_MessageData_ofExpr(x_236); +x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_358 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_358, 0, x_356); -lean_ctor_set(x_358, 1, x_357); -x_359 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_356); +x_359 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_360 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_360, 0, x_358); lean_ctor_set(x_360, 1, x_359); -x_361 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_361 = l_Lean_MessageData_ofExpr(x_246); x_362 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_362, 0, x_360); lean_ctor_set(x_362, 1, x_361); +x_363 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_364 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set(x_364, 1, x_363); lean_inc(x_12); -x_363 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_362, x_12, x_13, x_14, x_15, x_16, x_17, x_245); -x_247 = x_363; -goto block_277; +x_365 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_247); +x_249 = x_365; +goto block_279; } -block_277: +block_279: { -if (lean_obj_tag(x_247) == 0) +if (lean_obj_tag(x_249) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_250 = x_247; -} else { - lean_dec_ref(x_247); - x_250 = lean_box(0); -} -x_251 = lean_ctor_get(x_248, 0); +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); lean_inc(x_251); -x_252 = lean_ctor_get(x_248, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_253 = x_248; -} else { - lean_dec_ref(x_248); - x_253 = lean_box(0); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_252 = x_249; +} else { + lean_dec_ref(x_249); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_250, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_255 = x_250; +} else { + lean_dec_ref(x_250); + x_255 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_253)) { - x_254 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(0, 2, 0); } else { - x_254 = x_253; + x_256 = x_255; } -lean_ctor_set(x_254, 0, x_251); -lean_ctor_set(x_254, 1, x_6); -x_255 = lean_array_get_size(x_252); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_nat_dec_lt(x_256, x_255); -if (x_257 == 0) +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_6); +x_257 = lean_array_get_size(x_254); +x_258 = lean_unsigned_to_nat(0u); +x_259 = lean_nat_dec_lt(x_258, x_257); +if (x_259 == 0) { -lean_object* x_258; -lean_dec(x_255); +lean_object* x_260; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10332,25 +10335,25 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_258 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_260 = lean_alloc_ctor(0, 2, 0); } else { - x_258 = x_250; + x_260 = x_252; } -lean_ctor_set(x_258, 0, x_6); -lean_ctor_set(x_258, 1, x_249); -return x_258; +lean_ctor_set(x_260, 0, x_6); +lean_ctor_set(x_260, 1, x_251); +return x_260; } else { -uint8_t x_259; -x_259 = lean_nat_dec_le(x_255, x_255); -if (x_259 == 0) +uint8_t x_261; +x_261 = lean_nat_dec_le(x_257, x_257); +if (x_261 == 0) { -lean_object* x_260; -lean_dec(x_255); +lean_object* x_262; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10360,82 +10363,82 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_260 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_262 = lean_alloc_ctor(0, 2, 0); } else { - x_260 = x_250; + x_262 = x_252; } -lean_ctor_set(x_260, 0, x_6); -lean_ctor_set(x_260, 1, x_249); -return x_260; +lean_ctor_set(x_262, 0, x_6); +lean_ctor_set(x_262, 1, x_251); +return x_262; } else { -size_t x_261; size_t x_262; lean_object* x_263; -lean_dec(x_250); -lean_dec(x_6); -x_261 = 0; -x_262 = lean_usize_of_nat(x_255); -lean_dec(x_255); -x_263 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_252, x_261, x_262, x_254, x_12, x_13, x_14, x_15, x_16, x_17, x_249); +size_t x_263; size_t x_264; lean_object* x_265; lean_dec(x_252); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_266 = x_263; -} else { - lean_dec_ref(x_263); - x_266 = lean_box(0); -} -x_267 = lean_ctor_get(x_264, 1); +lean_dec(x_6); +x_263 = 0; +x_264 = lean_usize_of_nat(x_257); +lean_dec(x_257); +x_265 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__2___rarg(x_4, x_7, x_8, x_254, x_263, x_264, x_256, x_12, x_13, x_14, x_15, x_16, x_17, x_251); +lean_dec(x_254); +if (lean_obj_tag(x_265) == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_265, 1); lean_inc(x_267); -lean_dec(x_264); -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_268 = x_265; } else { - x_268 = x_266; + lean_dec_ref(x_265); + x_268 = lean_box(0); } -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_265); -return x_268; +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +lean_dec(x_266); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(0, 2, 0); +} else { + x_270 = x_268; +} +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_267); +return x_270; } else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_269 = lean_ctor_get(x_263, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_263, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_271 = x_263; +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_271 = lean_ctor_get(x_265, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_265, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_273 = x_265; } else { - lean_dec_ref(x_263); - x_271 = lean_box(0); + lean_dec_ref(x_265); + x_273 = lean_box(0); } -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_273)) { + x_274 = lean_alloc_ctor(1, 2, 0); } else { - x_272 = x_271; + x_274 = x_273; } -lean_ctor_set(x_272, 0, x_269); -lean_ctor_set(x_272, 1, x_270); -return x_272; +lean_ctor_set(x_274, 0, x_271); +lean_ctor_set(x_274, 1, x_272); +return x_274; } } } } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10446,33 +10449,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_273 = lean_ctor_get(x_247, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_247, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_275 = x_247; +x_275 = lean_ctor_get(x_249, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_249, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_277 = x_249; } else { - lean_dec_ref(x_247); - x_275 = lean_box(0); + lean_dec_ref(x_249); + x_277 = lean_box(0); } -if (lean_is_scalar(x_275)) { - x_276 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_276 = x_275; + x_278 = x_277; } -lean_ctor_set(x_276, 0, x_273); -lean_ctor_set(x_276, 1, x_274); -return x_276; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -lean_dec(x_235); +lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10488,32 +10491,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_364 = lean_ctor_get(x_243, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_243, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_243)) { - lean_ctor_release(x_243, 0); - lean_ctor_release(x_243, 1); - x_366 = x_243; +x_366 = lean_ctor_get(x_245, 0); +lean_inc(x_366); +x_367 = lean_ctor_get(x_245, 1); +lean_inc(x_367); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_368 = x_245; } else { - lean_dec_ref(x_243); - x_366 = lean_box(0); + lean_dec_ref(x_245); + x_368 = lean_box(0); } -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_368)) { + x_369 = lean_alloc_ctor(1, 2, 0); } else { - x_367 = x_366; + x_369 = x_368; } -lean_ctor_set(x_367, 0, x_364); -lean_ctor_set(x_367, 1, x_365); -return x_367; +lean_ctor_set(x_369, 0, x_366); +lean_ctor_set(x_369, 1, x_367); +return x_369; } } else { -lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; -lean_dec(x_235); +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10529,32 +10532,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_368 = lean_ctor_get(x_240, 0); -lean_inc(x_368); -x_369 = lean_ctor_get(x_240, 1); -lean_inc(x_369); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_370 = x_240; +x_370 = lean_ctor_get(x_242, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_242, 1); +lean_inc(x_371); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_372 = x_242; } else { - lean_dec_ref(x_240); - x_370 = lean_box(0); + lean_dec_ref(x_242); + x_372 = lean_box(0); } -if (lean_is_scalar(x_370)) { - x_371 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_372)) { + x_373 = lean_alloc_ctor(1, 2, 0); } else { - x_371 = x_370; + x_373 = x_372; } -lean_ctor_set(x_371, 0, x_368); -lean_ctor_set(x_371, 1, x_369); -return x_371; +lean_ctor_set(x_373, 0, x_370); +lean_ctor_set(x_373, 1, x_371); +return x_373; } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_235); +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10570,32 +10573,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_372 = lean_ctor_get(x_238, 0); -lean_inc(x_372); -x_373 = lean_ctor_get(x_238, 1); -lean_inc(x_373); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - lean_ctor_release(x_238, 1); - x_374 = x_238; +x_374 = lean_ctor_get(x_240, 0); +lean_inc(x_374); +x_375 = lean_ctor_get(x_240, 1); +lean_inc(x_375); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_376 = x_240; } else { - lean_dec_ref(x_238); - x_374 = lean_box(0); + lean_dec_ref(x_240); + x_376 = lean_box(0); } -if (lean_is_scalar(x_374)) { - x_375 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_376)) { + x_377 = lean_alloc_ctor(1, 2, 0); } else { - x_375 = x_374; + x_377 = x_376; } -lean_ctor_set(x_375, 0, x_372); -lean_ctor_set(x_375, 1, x_373); -return x_375; +lean_ctor_set(x_377, 0, x_374); +lean_ctor_set(x_377, 1, x_375); +return x_377; } } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -lean_dec(x_235); +lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10611,33 +10614,33 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_376 = lean_ctor_get(x_236, 0); -lean_inc(x_376); -x_377 = lean_ctor_get(x_236, 1); -lean_inc(x_377); -if (lean_is_exclusive(x_236)) { - lean_ctor_release(x_236, 0); - lean_ctor_release(x_236, 1); - x_378 = x_236; +x_378 = lean_ctor_get(x_237, 0); +lean_inc(x_378); +x_379 = lean_ctor_get(x_237, 1); +lean_inc(x_379); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_380 = x_237; } else { - lean_dec_ref(x_236); - x_378 = lean_box(0); + lean_dec_ref(x_237); + x_380 = lean_box(0); } -if (lean_is_scalar(x_378)) { - x_379 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_380)) { + x_381 = lean_alloc_ctor(1, 2, 0); } else { - x_379 = x_378; + x_381 = x_380; } -lean_ctor_set(x_379, 0, x_376); -lean_ctor_set(x_379, 1, x_377); -return x_379; +lean_ctor_set(x_381, 0, x_378); +lean_ctor_set(x_381, 1, x_379); +return x_381; } } } } else { -uint8_t x_389; +uint8_t x_391; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -10654,23 +10657,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_389 = !lean_is_exclusive(x_21); -if (x_389 == 0) +x_391 = !lean_is_exclusive(x_21); +if (x_391 == 0) { return x_21; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_21, 0); -x_391 = lean_ctor_get(x_21, 1); -lean_inc(x_391); -lean_inc(x_390); +lean_object* x_392; lean_object* x_393; lean_object* x_394; +x_392 = lean_ctor_get(x_21, 0); +x_393 = lean_ctor_get(x_21, 1); +lean_inc(x_393); +lean_inc(x_392); lean_dec(x_21); -x_392 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_392, 0, x_390); -lean_ctor_set(x_392, 1, x_391); -return x_392; +x_394 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_394, 0, x_392); +lean_ctor_set(x_394, 1, x_393); +return x_394; } } } @@ -11062,148 +11065,149 @@ lean_dec(x_24); x_25 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); if (lean_obj_tag(x_25) == 1) { -lean_object* x_227; -x_227 = lean_ctor_get(x_25, 0); -lean_inc(x_227); -lean_ctor_set(x_21, 0, x_227); +lean_object* x_228; +x_228 = lean_ctor_get(x_25, 0); +lean_inc(x_228); +lean_ctor_set(x_21, 0, x_228); x_26 = x_21; -goto block_226; +goto block_227; } else { -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_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_free_object(x_21); lean_inc(x_25); -x_228 = l_Lean_MessageData_ofExpr(x_25); -x_229 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_230 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_228); -x_231 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; -x_232 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_232, 0, x_230); -lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_232, x_14, x_15, x_16, x_17, x_23); -x_26 = x_233; -goto block_226; -} -block_226: +x_229 = l_Lean_MessageData_ofExpr(x_25); +x_230 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_231 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_229); +x_232 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_233 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_233, x_14, x_15, x_16, x_17, x_23); +x_26 = x_234; +goto block_227; +} +block_227: { if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_27; uint8_t x_28; lean_object* x_29; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); +x_28 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_28 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -if (lean_obj_tag(x_28) == 0) +x_29 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_28, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_25); -x_30 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_33 = l_Lean_Meta_whnfD(x_31, x_14, x_15, x_16, x_17, x_32); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_Meta_whnfD(x_32, x_14, x_15, x_16, x_17, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Expr_getAppFn(x_34); -if (lean_obj_tag(x_36) == 4) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Expr_getAppFn(x_35); +if (lean_obj_tag(x_37) == 4) { -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_ctor_get(x_36, 0); -lean_inc(x_110); -lean_dec(x_36); -x_111 = lean_st_ref_get(x_17, x_35); -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_111 = lean_ctor_get(x_37, 0); +lean_inc(x_111); +lean_dec(x_37); +x_112 = lean_st_ref_get(x_17, x_36); +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -lean_dec(x_111); -x_114 = lean_ctor_get(x_112, 0); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_110); -if (lean_obj_tag(x_115) == 0) +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_environment_find(x_115, x_111); +if (lean_obj_tag(x_116) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_object* x_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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_116 = l_Lean_MessageData_ofExpr(x_25); -x_117 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_116); -x_119 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_MessageData_ofExpr(x_34); -x_122 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_124 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); +x_117 = l_Lean_MessageData_ofExpr(x_25); +x_118 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_119 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_121 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Lean_MessageData_ofExpr(x_35); +x_123 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_125 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); lean_inc(x_12); -x_125 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_124, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_125; -goto block_109; +x_126 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_125, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_126; +goto block_110; } else { -lean_object* x_126; -x_126 = lean_ctor_get(x_115, 0); -lean_inc(x_126); -lean_dec(x_115); -switch (lean_obj_tag(x_126)) { +lean_object* x_127; +x_127 = lean_ctor_get(x_116, 0); +lean_inc(x_127); +lean_dec(x_116); +switch (lean_obj_tag(x_127)) { case 4: { -lean_object* x_127; lean_object* x_128; uint8_t x_144; lean_object* x_145; +lean_object* x_128; lean_object* x_129; uint8_t x_145; lean_object* x_146; lean_dec(x_1); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -lean_dec(x_126); -x_144 = lean_ctor_get_uint8(x_127, sizeof(void*)*1); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); lean_dec(x_127); -x_145 = lean_box(x_144); -if (lean_obj_tag(x_145) == 0) +x_145 = lean_ctor_get_uint8(x_128, sizeof(void*)*1); +lean_dec(x_128); +x_146 = lean_box(x_145); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_34); -x_146 = lean_box(0); +lean_object* x_147; lean_object* x_148; +lean_dec(x_35); +x_147 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -11211,357 +11215,357 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_147 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_146, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_147; -goto block_109; +x_148 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_148; +goto block_110; } else { -lean_object* x_148; -lean_dec(x_145); +lean_object* x_149; +lean_dec(x_146); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_148 = lean_box(0); -x_128 = x_148; -goto block_143; +x_149 = lean_box(0); +x_129 = x_149; +goto block_144; } -block_143: +block_144: { -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; uint8_t x_139; -lean_dec(x_128); -x_129 = l_Lean_MessageData_ofExpr(x_25); -x_130 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_131 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_133 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_MessageData_ofExpr(x_34); -x_135 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_137 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); +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; uint8_t x_140; +lean_dec(x_129); +x_130 = l_Lean_MessageData_ofExpr(x_25); +x_131 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_132 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_130); +x_133 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_134 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_MessageData_ofExpr(x_35); +x_136 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_138 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); lean_inc(x_12); -x_138 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_137, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_139 = !lean_is_exclusive(x_138); -if (x_139 == 0) +x_139 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_138, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) { -x_37 = x_138; -goto block_109; +x_38 = x_139; +goto block_110; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_138, 0); -x_141 = lean_ctor_get(x_138, 1); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_139, 0); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_138); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_37 = x_142; -goto block_109; +lean_dec(x_139); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_38 = x_143; +goto block_110; } } } case 5: { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_34); +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_35); lean_dec(x_11); lean_dec(x_3); -x_149 = lean_ctor_get(x_126, 0); -lean_inc(x_149); -lean_dec(x_126); -x_150 = lean_ctor_get(x_149, 1); +x_150 = lean_ctor_get(x_127, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 4); +lean_dec(x_127); +x_151 = lean_ctor_get(x_150, 1); lean_inc(x_151); -lean_dec(x_149); -x_152 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_153 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +x_152 = lean_ctor_get(x_150, 4); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_154 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_154 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_150, x_153, x_151, x_152, x_14, x_15, x_16, x_17, x_113); -lean_dec(x_150); -if (lean_obj_tag(x_154) == 0) +x_155 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_151, x_154, x_152, x_153, x_14, x_15, x_16, x_17, x_114); +lean_dec(x_151); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); -lean_dec(x_154); -x_157 = !lean_is_exclusive(x_155); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; -x_158 = lean_ctor_get(x_155, 0); -x_159 = lean_ctor_get(x_155, 1); -x_160 = l_Lean_Expr_fvarId_x21(x_25); -x_161 = 1; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_158 = !lean_is_exclusive(x_156); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; +x_159 = lean_ctor_get(x_156, 0); +x_160 = lean_ctor_get(x_156, 1); +x_161 = l_Lean_Expr_fvarId_x21(x_25); +x_162 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_162 = l_Lean_Meta_Cases_cases(x_10, x_160, x_158, x_161, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_162) == 0) +x_163 = l_Lean_Meta_Cases_cases(x_10, x_161, x_159, x_162, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_163) == 0) { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_162); -if (x_163 == 0) +uint8_t x_164; +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) { -lean_object* x_164; -x_164 = lean_ctor_get(x_162, 0); -lean_ctor_set(x_155, 1, x_164); -lean_ctor_set(x_155, 0, x_159); -lean_ctor_set(x_162, 0, x_155); -x_37 = x_162; -goto block_109; +lean_object* x_165; +x_165 = lean_ctor_get(x_163, 0); +lean_ctor_set(x_156, 1, x_165); +lean_ctor_set(x_156, 0, x_160); +lean_ctor_set(x_163, 0, x_156); +x_38 = x_163; +goto block_110; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_162, 0); -x_166 = lean_ctor_get(x_162, 1); +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_163, 0); +x_167 = lean_ctor_get(x_163, 1); +lean_inc(x_167); lean_inc(x_166); -lean_inc(x_165); -lean_dec(x_162); -lean_ctor_set(x_155, 1, x_165); -lean_ctor_set(x_155, 0, x_159); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_155); -lean_ctor_set(x_167, 1, x_166); -x_37 = x_167; -goto block_109; +lean_dec(x_163); +lean_ctor_set(x_156, 1, x_166); +lean_ctor_set(x_156, 0, x_160); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_156); +lean_ctor_set(x_168, 1, x_167); +x_38 = x_168; +goto block_110; } } else { -uint8_t x_168; -lean_free_object(x_155); -lean_dec(x_159); -x_168 = !lean_is_exclusive(x_162); -if (x_168 == 0) +uint8_t x_169; +lean_free_object(x_156); +lean_dec(x_160); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -x_37 = x_162; -goto block_109; +x_38 = x_163; +goto block_110; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_162, 0); -x_170 = lean_ctor_get(x_162, 1); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_163, 0); +x_171 = lean_ctor_get(x_163, 1); +lean_inc(x_171); lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_162); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_37 = x_171; -goto block_109; +lean_dec(x_163); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_38 = x_172; +goto block_110; } } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; -x_172 = lean_ctor_get(x_155, 0); -x_173 = lean_ctor_get(x_155, 1); +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; +x_173 = lean_ctor_get(x_156, 0); +x_174 = lean_ctor_get(x_156, 1); +lean_inc(x_174); lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_155); -x_174 = l_Lean_Expr_fvarId_x21(x_25); -x_175 = 1; +lean_dec(x_156); +x_175 = l_Lean_Expr_fvarId_x21(x_25); +x_176 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_176 = l_Lean_Meta_Cases_cases(x_10, x_174, x_172, x_175, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_176) == 0) +x_177 = l_Lean_Meta_Cases_cases(x_10, x_175, x_173, x_176, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_177) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_179 = x_176; +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_180 = x_177; } else { - lean_dec_ref(x_176); - x_179 = lean_box(0); + lean_dec_ref(x_177); + x_180 = lean_box(0); } -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_173); -lean_ctor_set(x_180, 1, x_177); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(0, 2, 0); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_174); +lean_ctor_set(x_181, 1, x_178); +if (lean_is_scalar(x_180)) { + x_182 = lean_alloc_ctor(0, 2, 0); } else { - x_181 = x_179; + x_182 = x_180; } -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_178); -x_37 = x_181; -goto block_109; +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +x_38 = x_182; +goto block_110; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_173); -x_182 = lean_ctor_get(x_176, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_176, 1); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_174); +x_183 = lean_ctor_get(x_177, 0); lean_inc(x_183); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_184 = x_176; +x_184 = lean_ctor_get(x_177, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_185 = x_177; } else { - lean_dec_ref(x_176); - x_184 = lean_box(0); + lean_dec_ref(x_177); + x_185 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_186 = x_185; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -x_37 = x_185; -goto block_109; +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_184); +x_38 = x_186; +goto block_110; } } } else { -uint8_t x_186; +uint8_t x_187; lean_dec(x_25); lean_dec(x_10); -x_186 = !lean_is_exclusive(x_154); -if (x_186 == 0) +x_187 = !lean_is_exclusive(x_155); +if (x_187 == 0) { -x_37 = x_154; -goto block_109; +x_38 = x_155; +goto block_110; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_154, 0); -x_188 = lean_ctor_get(x_154, 1); +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_155, 0); +x_189 = lean_ctor_get(x_155, 1); +lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_154); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -x_37 = x_189; -goto block_109; +lean_dec(x_155); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_38 = x_190; +goto block_110; } } } default: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_126); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_127); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_190 = l_Lean_MessageData_ofExpr(x_25); -x_191 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_192 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_190); -x_193 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_MessageData_ofExpr(x_34); -x_196 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -x_197 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_198 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); +x_191 = l_Lean_MessageData_ofExpr(x_25); +x_192 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_MessageData_ofExpr(x_35); +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 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_199 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); lean_inc(x_12); -x_199 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_199; -goto block_109; +x_200 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_199, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_200; +goto block_110; } } } } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_36); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_37); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_200 = l_Lean_MessageData_ofExpr(x_25); -x_201 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_202 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_200); -x_203 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_204 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -x_205 = l_Lean_MessageData_ofExpr(x_34); -x_206 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_208 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); +x_201 = l_Lean_MessageData_ofExpr(x_25); +x_202 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_203 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_205 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_MessageData_ofExpr(x_35); +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +x_208 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_209 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); lean_inc(x_12); -x_209 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_208, x_12, x_13, x_14, x_15, x_16, x_17, x_35); -x_37 = x_209; -goto block_109; +x_210 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_209, x_12, x_13, x_14, x_15, x_16, x_17, x_36); +x_38 = x_210; +goto block_110; } -block_109: +block_110: { -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_38) == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_38, 0); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_37, 1); -x_42 = lean_ctor_get(x_39, 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_38, 1); +x_43 = lean_ctor_get(x_40, 1); lean_inc(x_6); -lean_ctor_set(x_39, 1, x_6); -x_43 = lean_array_get_size(x_42); -x_44 = lean_unsigned_to_nat(0u); -x_45 = lean_nat_dec_lt(x_44, x_43); -if (x_45 == 0) +lean_ctor_set(x_40, 1, x_6); +x_44 = lean_array_get_size(x_43); +x_45 = lean_unsigned_to_nat(0u); +x_46 = lean_nat_dec_lt(x_45, x_44); +if (x_46 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11571,18 +11575,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_43, x_43); -if (x_46 == 0) +uint8_t x_47; +x_47 = lean_nat_dec_le(x_44, x_44); +if (x_47 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11592,70 +11596,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_47; size_t x_48; lean_object* x_49; -lean_free_object(x_37); +size_t x_48; size_t x_49; lean_object* x_50; +lean_free_object(x_38); lean_dec(x_6); -x_47 = 0; -x_48 = lean_usize_of_nat(x_43); +x_48 = 0; +x_49 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_43, x_48, x_49, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_42); lean_dec(x_43); -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_42, x_47, x_48, x_39, x_12, x_13, x_14, x_15, x_16, x_17, x_41); -lean_dec(x_42); -if (lean_obj_tag(x_49) == 0) +if (lean_obj_tag(x_50) == 0) { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_52); -return x_49; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_53); +return x_50; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_49, 0); -x_54 = lean_ctor_get(x_49, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_49); -x_55 = lean_ctor_get(x_53, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +x_55 = lean_ctor_get(x_50, 1); lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_inc(x_54); +lean_dec(x_50); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_58; +x_58 = !lean_is_exclusive(x_50); +if (x_58 == 0) { -return x_49; +return x_50; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_50, 0); +x_60 = lean_ctor_get(x_50, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_50); +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; } } } @@ -11663,25 +11667,25 @@ return x_60; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_61 = lean_ctor_get(x_37, 1); -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_62 = lean_ctor_get(x_38, 1); +x_63 = lean_ctor_get(x_40, 0); +x_64 = lean_ctor_get(x_40, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); +lean_dec(x_40); lean_inc(x_6); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_6); -x_65 = lean_array_get_size(x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_6); +x_66 = lean_array_get_size(x_64); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +if (x_68 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11691,18 +11695,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_68; -x_68 = lean_nat_dec_le(x_65, x_65); -if (x_68 == 0) +uint8_t x_69; +x_69 = lean_nat_dec_le(x_66, x_66); +if (x_69 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11712,69 +11716,69 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_69; size_t x_70; lean_object* x_71; -lean_free_object(x_37); +size_t x_70; size_t x_71; lean_object* x_72; +lean_free_object(x_38); lean_dec(x_6); -x_69 = 0; -x_70 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_71 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_63, x_69, x_70, x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -lean_dec(x_63); -if (lean_obj_tag(x_71) == 0) +x_70 = 0; +x_71 = lean_usize_of_nat(x_66); +lean_dec(x_66); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_64, x_70, x_71, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_64); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); 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_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_72); + x_75 = lean_box(0); } -x_75 = lean_ctor_get(x_72, 1); -lean_inc(x_75); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_76 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_76 = x_74; + x_77 = x_75; } -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_73); -return x_76; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_71, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_71, 1); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_72, 0); lean_inc(x_78); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_79 = x_71; +x_79 = lean_ctor_get(x_72, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_80 = x_72; } else { - lean_dec_ref(x_71); - x_79 = lean_box(0); + lean_dec_ref(x_72); + x_80 = lean_box(0); } -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_80 = x_79; + x_81 = x_80; } -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_78); -return x_80; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } @@ -11782,41 +11786,41 @@ 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; uint8_t x_89; -x_81 = lean_ctor_get(x_37, 0); -x_82 = lean_ctor_get(x_37, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_37); -x_83 = lean_ctor_get(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; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_82 = lean_ctor_get(x_38, 0); +x_83 = lean_ctor_get(x_38, 1); lean_inc(x_83); -x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_38); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_85 = x_81; +x_85 = lean_ctor_get(x_82, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_81); - x_85 = lean_box(0); + lean_dec_ref(x_82); + x_86 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_86 = x_85; + x_87 = x_86; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_6); -x_87 = lean_array_get_size(x_84); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_nat_dec_lt(x_88, x_87); -if (x_89 == 0) +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_6); +x_88 = lean_array_get_size(x_85); +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_nat_dec_lt(x_89, x_88); +if (x_90 == 0) { -lean_object* x_90; +lean_object* x_91; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11826,21 +11830,21 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_82); -return x_90; +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_6); +lean_ctor_set(x_91, 1, x_83); +return x_91; } else { -uint8_t x_91; -x_91 = lean_nat_dec_le(x_87, x_87); -if (x_91 == 0) +uint8_t x_92; +x_92 = lean_nat_dec_le(x_88, x_88); +if (x_92 == 0) { -lean_object* x_92; +lean_object* x_93; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11850,70 +11854,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_6); -lean_ctor_set(x_92, 1, x_82); -return x_92; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_6); +lean_ctor_set(x_93, 1, x_83); +return x_93; } else { -size_t x_93; size_t x_94; lean_object* x_95; +size_t x_94; size_t x_95; lean_object* x_96; lean_dec(x_6); -x_93 = 0; -x_94 = lean_usize_of_nat(x_87); -lean_dec(x_87); -x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_84, x_93, x_94, x_86, x_12, x_13, x_14, x_15, x_16, x_17, x_82); -lean_dec(x_84); -if (lean_obj_tag(x_95) == 0) +x_94 = 0; +x_95 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_96 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_85, x_94, x_95, x_87, x_12, x_13, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_85); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_ctor_get(x_96, 0); lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_99 = x_96; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_96); + x_99 = lean_box(0); } -x_99 = lean_ctor_get(x_96, 1); -lean_inc(x_99); -lean_dec(x_96); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_101 = x_99; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_101 = lean_ctor_get(x_95, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_95, 1); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_96, 0); lean_inc(x_102); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_103 = x_95; +x_103 = lean_ctor_get(x_96, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_104 = x_96; } else { - lean_dec_ref(x_95); - x_103 = lean_box(0); + lean_dec_ref(x_96); + x_104 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_105 = x_104; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } } @@ -11921,7 +11925,7 @@ return x_104; } else { -uint8_t x_105; +uint8_t x_106; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -11932,30 +11936,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_37); -if (x_105 == 0) +x_106 = !lean_is_exclusive(x_38); +if (x_106 == 0) { -return x_37; +return x_38; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_37, 0); -x_107 = lean_ctor_get(x_37, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_38, 0); +x_108 = lean_ctor_get(x_38, 1); +lean_inc(x_108); lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_37); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_dec(x_38); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; } } } } else { -uint8_t x_210; +uint8_t x_211; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -11972,29 +11976,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_210 = !lean_is_exclusive(x_33); -if (x_210 == 0) +x_211 = !lean_is_exclusive(x_34); +if (x_211 == 0) { -return x_33; +return x_34; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_33, 0); -x_212 = lean_ctor_get(x_33, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_34, 0); +x_213 = lean_ctor_get(x_34, 1); +lean_inc(x_213); lean_inc(x_212); -lean_inc(x_211); -lean_dec(x_33); -x_213 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -return x_213; +lean_dec(x_34); +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +return x_214; } } } else { -uint8_t x_214; +uint8_t x_215; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -12011,29 +12015,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_214 = !lean_is_exclusive(x_30); -if (x_214 == 0) +x_215 = !lean_is_exclusive(x_31); +if (x_215 == 0) { -return x_30; +return x_31; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_215 = lean_ctor_get(x_30, 0); -x_216 = lean_ctor_get(x_30, 1); +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_ctor_get(x_31, 0); +x_217 = lean_ctor_get(x_31, 1); +lean_inc(x_217); lean_inc(x_216); -lean_inc(x_215); -lean_dec(x_30); -x_217 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -return x_217; +lean_dec(x_31); +x_218 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +return x_218; } } } else { -uint8_t x_218; +uint8_t x_219; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -12050,29 +12054,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_218 = !lean_is_exclusive(x_28); -if (x_218 == 0) +x_219 = !lean_is_exclusive(x_29); +if (x_219 == 0) { -return x_28; +return x_29; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_219 = lean_ctor_get(x_28, 0); -x_220 = lean_ctor_get(x_28, 1); +lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_220 = lean_ctor_get(x_29, 0); +x_221 = lean_ctor_get(x_29, 1); +lean_inc(x_221); lean_inc(x_220); -lean_inc(x_219); -lean_dec(x_28); -x_221 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -return x_221; +lean_dec(x_29); +x_222 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } } else { -uint8_t x_222; +uint8_t x_223; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -12089,179 +12093,180 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_222 = !lean_is_exclusive(x_26); -if (x_222 == 0) +x_223 = !lean_is_exclusive(x_26); +if (x_223 == 0) { return x_26; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_26, 0); -x_224 = lean_ctor_get(x_26, 1); +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_26, 0); +x_225 = lean_ctor_get(x_26, 1); +lean_inc(x_225); lean_inc(x_224); -lean_inc(x_223); lean_dec(x_26); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +x_226 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_226, 0, x_224); +lean_ctor_set(x_226, 1, x_225); +return x_226; } } } } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_234 = lean_ctor_get(x_21, 1); -lean_inc(x_234); +lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_235 = lean_ctor_get(x_21, 1); +lean_inc(x_235); lean_dec(x_21); -x_235 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); -if (lean_obj_tag(x_235) == 1) +x_236 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); +if (lean_obj_tag(x_236) == 1) { -lean_object* x_381; lean_object* x_382; -x_381 = lean_ctor_get(x_235, 0); -lean_inc(x_381); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_381); -lean_ctor_set(x_382, 1, x_234); -x_236 = x_382; -goto block_380; +lean_object* x_383; lean_object* x_384; +x_383 = lean_ctor_get(x_236, 0); +lean_inc(x_383); +x_384 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_235); +x_237 = x_384; +goto block_382; } else { -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_inc(x_235); -x_383 = l_Lean_MessageData_ofExpr(x_235); -x_384 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_385 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_383); -x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +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_inc(x_236); +x_385 = l_Lean_MessageData_ofExpr(x_236); +x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_387 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_387, 0, x_385); -lean_ctor_set(x_387, 1, x_386); -x_388 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_387, x_14, x_15, x_16, x_17, x_234); -x_236 = x_388; -goto block_380; -} -block_380: -{ -if (lean_obj_tag(x_236) == 0) -{ -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_236, 1); -lean_inc(x_237); -lean_dec(x_236); +lean_ctor_set(x_387, 0, x_386); +lean_ctor_set(x_387, 1, x_385); +x_388 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_389 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_389, 0, x_387); +lean_ctor_set(x_389, 1, x_388); +x_390 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_389, x_14, x_15, x_16, x_17, x_235); +x_237 = x_390; +goto block_382; +} +block_382: +{ +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; uint8_t x_239; lean_object* x_240; +x_238 = lean_ctor_get(x_237, 1); +lean_inc(x_238); +lean_dec(x_237); +x_239 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_238 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_237); -if (lean_obj_tag(x_238) == 0) -{ -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_238, 1); -lean_inc(x_239); -lean_dec(x_238); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_235); -x_240 = lean_infer_type(x_235, x_14, x_15, x_16, x_17, x_239); +x_240 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_239, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_238); if (lean_obj_tag(x_240) == 0) { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_240, 0); +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_240, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); lean_dec(x_240); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_243 = l_Lean_Meta_whnfD(x_241, x_14, x_15, x_16, x_17, x_242); -if (lean_obj_tag(x_243) == 0) +lean_inc(x_236); +x_242 = lean_infer_type(x_236, x_14, x_15, x_16, x_17, x_241); +if (lean_obj_tag(x_242) == 0) { -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_244 = lean_ctor_get(x_243, 0); +lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); -x_245 = lean_ctor_get(x_243, 1); -lean_inc(x_245); -lean_dec(x_243); -x_246 = l_Lean_Expr_getAppFn(x_244); -if (lean_obj_tag(x_246) == 4) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_278 = lean_ctor_get(x_246, 0); -lean_inc(x_278); -lean_dec(x_246); -x_279 = lean_st_ref_get(x_17, x_245); -x_280 = lean_ctor_get(x_279, 0); +lean_dec(x_242); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_245 = l_Lean_Meta_whnfD(x_243, x_14, x_15, x_16, x_17, x_244); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec(x_245); +x_248 = l_Lean_Expr_getAppFn(x_246); +if (lean_obj_tag(x_248) == 4) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_280 = lean_ctor_get(x_248, 0); lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_282 = lean_ctor_get(x_280, 0); +lean_dec(x_248); +x_281 = lean_st_ref_get(x_17, x_247); +x_282 = lean_ctor_get(x_281, 0); lean_inc(x_282); -lean_dec(x_280); -x_283 = lean_environment_find(x_282, x_278); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_ctor_get(x_282, 0); +lean_inc(x_284); +lean_dec(x_282); +x_285 = lean_environment_find(x_284, x_280); +if (lean_obj_tag(x_285) == 0) +{ +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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_284 = l_Lean_MessageData_ofExpr(x_235); -x_285 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_286 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_284); -x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_286 = l_Lean_MessageData_ofExpr(x_236); +x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_288 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_288, 0, x_286); -lean_ctor_set(x_288, 1, x_287); -x_289 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_286); +x_289 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_290 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_290, 0, x_288); lean_ctor_set(x_290, 1, x_289); -x_291 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_291 = l_Lean_MessageData_ofExpr(x_246); x_292 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_292, 0, x_290); lean_ctor_set(x_292, 1, x_291); +x_293 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_294 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); lean_inc(x_12); -x_293 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_293; -goto block_277; +x_295 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_294, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_295; +goto block_279; } else { -lean_object* x_294; -x_294 = lean_ctor_get(x_283, 0); -lean_inc(x_294); -lean_dec(x_283); -switch (lean_obj_tag(x_294)) { +lean_object* x_296; +x_296 = lean_ctor_get(x_285, 0); +lean_inc(x_296); +lean_dec(x_285); +switch (lean_obj_tag(x_296)) { case 4: { -lean_object* x_295; lean_object* x_296; uint8_t x_312; lean_object* x_313; +lean_object* x_297; lean_object* x_298; uint8_t x_314; lean_object* x_315; lean_dec(x_1); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -lean_dec(x_294); -x_312 = lean_ctor_get_uint8(x_295, sizeof(void*)*1); -lean_dec(x_295); -x_313 = lean_box(x_312); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; lean_object* x_315; -lean_dec(x_244); -x_314 = lean_box(0); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +lean_dec(x_296); +x_314 = lean_ctor_get_uint8(x_297, sizeof(void*)*1); +lean_dec(x_297); +x_315 = lean_box(x_314); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; +lean_dec(x_246); +x_316 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -12269,316 +12274,316 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_315 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_235, x_10, x_3, x_11, x_4, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_315; -goto block_277; +x_317 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_236, x_10, x_3, x_11, x_4, x_316, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_317; +goto block_279; } else { -lean_object* x_316; -lean_dec(x_313); +lean_object* x_318; +lean_dec(x_315); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_316 = lean_box(0); -x_296 = x_316; -goto block_311; +x_318 = lean_box(0); +x_298 = x_318; +goto block_313; } -block_311: +block_313: { -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_296); -x_297 = l_Lean_MessageData_ofExpr(x_235); -x_298 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_299 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +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_dec(x_298); +x_299 = l_Lean_MessageData_ofExpr(x_236); +x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_301 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_299); +x_302 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_303 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_303, 0, x_301); lean_ctor_set(x_303, 1, x_302); -x_304 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_304 = l_Lean_MessageData_ofExpr(x_246); x_305 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_305, 0, x_303); lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_307 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_307, 0, x_305); +lean_ctor_set(x_307, 1, x_306); lean_inc(x_12); -x_306 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_305, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_307 = lean_ctor_get(x_306, 0); -lean_inc(x_307); -x_308 = lean_ctor_get(x_306, 1); -lean_inc(x_308); -if (lean_is_exclusive(x_306)) { - lean_ctor_release(x_306, 0); - lean_ctor_release(x_306, 1); - x_309 = x_306; +x_308 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_307, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_311 = x_308; } else { - lean_dec_ref(x_306); - x_309 = lean_box(0); + lean_dec_ref(x_308); + x_311 = lean_box(0); } -if (lean_is_scalar(x_309)) { - x_310 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_311)) { + x_312 = lean_alloc_ctor(1, 2, 0); } else { - x_310 = x_309; + x_312 = x_311; } -lean_ctor_set(x_310, 0, x_307); -lean_ctor_set(x_310, 1, x_308); -x_247 = x_310; -goto block_277; +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_310); +x_249 = x_312; +goto block_279; } } case 5: { -lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_244); +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_246); lean_dec(x_11); lean_dec(x_3); -x_317 = lean_ctor_get(x_294, 0); -lean_inc(x_317); -lean_dec(x_294); -x_318 = lean_ctor_get(x_317, 1); -lean_inc(x_318); -x_319 = lean_ctor_get(x_317, 4); +x_319 = lean_ctor_get(x_296, 0); lean_inc(x_319); -lean_dec(x_317); -x_320 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_321 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +lean_dec(x_296); +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +x_321 = lean_ctor_get(x_319, 4); +lean_inc(x_321); +lean_dec(x_319); +x_322 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_323 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_322 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_318, x_321, x_319, x_320, x_14, x_15, x_16, x_17, x_281); -lean_dec(x_318); -if (lean_obj_tag(x_322) == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; lean_object* x_330; -x_323 = lean_ctor_get(x_322, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_322, 1); -lean_inc(x_324); -lean_dec(x_322); -x_325 = lean_ctor_get(x_323, 0); +x_324 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_320, x_323, x_321, x_322, x_14, x_15, x_16, x_17, x_283); +lean_dec(x_320); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; lean_object* x_332; +x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); -x_326 = lean_ctor_get(x_323, 1); +x_326 = lean_ctor_get(x_324, 1); lean_inc(x_326); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_327 = x_323; -} else { - lean_dec_ref(x_323); - x_327 = lean_box(0); -} -x_328 = l_Lean_Expr_fvarId_x21(x_235); -x_329 = 1; +lean_dec(x_324); +x_327 = lean_ctor_get(x_325, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_329 = x_325; +} else { + lean_dec_ref(x_325); + x_329 = lean_box(0); +} +x_330 = l_Lean_Expr_fvarId_x21(x_236); +x_331 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_330 = l_Lean_Meta_Cases_cases(x_10, x_328, x_325, x_329, x_14, x_15, x_16, x_17, x_324); -if (lean_obj_tag(x_330) == 0) +x_332 = l_Lean_Meta_Cases_cases(x_10, x_330, x_327, x_331, x_14, x_15, x_16, x_17, x_326); +if (lean_obj_tag(x_332) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_330, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_333 = x_330; +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_332, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_335 = x_332; } else { - lean_dec_ref(x_330); - x_333 = lean_box(0); + lean_dec_ref(x_332); + x_335 = lean_box(0); } -if (lean_is_scalar(x_327)) { - x_334 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_329)) { + x_336 = lean_alloc_ctor(0, 2, 0); } else { - x_334 = x_327; + x_336 = x_329; } -lean_ctor_set(x_334, 0, x_326); -lean_ctor_set(x_334, 1, x_331); -if (lean_is_scalar(x_333)) { - x_335 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_328); +lean_ctor_set(x_336, 1, x_333); +if (lean_is_scalar(x_335)) { + x_337 = lean_alloc_ctor(0, 2, 0); } else { - x_335 = x_333; + x_337 = x_335; } -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_332); -x_247 = x_335; -goto block_277; +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_334); +x_249 = x_337; +goto block_279; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_327); -lean_dec(x_326); -x_336 = lean_ctor_get(x_330, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_330, 1); -lean_inc(x_337); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_338 = x_330; +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_329); +lean_dec(x_328); +x_338 = lean_ctor_get(x_332, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_332, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_340 = x_332; } else { - lean_dec_ref(x_330); - x_338 = lean_box(0); + lean_dec_ref(x_332); + x_340 = lean_box(0); } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); } else { - x_339 = x_338; + x_341 = x_340; } -lean_ctor_set(x_339, 0, x_336); -lean_ctor_set(x_339, 1, x_337); -x_247 = x_339; -goto block_277; +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +x_249 = x_341; +goto block_279; } } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_235); +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_236); lean_dec(x_10); -x_340 = lean_ctor_get(x_322, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_322, 1); -lean_inc(x_341); -if (lean_is_exclusive(x_322)) { - lean_ctor_release(x_322, 0); - lean_ctor_release(x_322, 1); - x_342 = x_322; +x_342 = lean_ctor_get(x_324, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_324, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_344 = x_324; } else { - lean_dec_ref(x_322); - x_342 = lean_box(0); + lean_dec_ref(x_324); + x_344 = lean_box(0); } -if (lean_is_scalar(x_342)) { - x_343 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); } else { - x_343 = x_342; + x_345 = x_344; } -lean_ctor_set(x_343, 0, x_340); -lean_ctor_set(x_343, 1, x_341); -x_247 = x_343; -goto block_277; +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +x_249 = x_345; +goto block_279; } } default: { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -lean_dec(x_294); +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_296); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_344 = l_Lean_MessageData_ofExpr(x_235); -x_345 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_346 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_344); -x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_346 = l_Lean_MessageData_ofExpr(x_236); +x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_348 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_348, 0, x_346); -lean_ctor_set(x_348, 1, x_347); -x_349 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_348, 0, x_347); +lean_ctor_set(x_348, 1, x_346); +x_349 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_350 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_350, 0, x_348); lean_ctor_set(x_350, 1, x_349); -x_351 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_351 = l_Lean_MessageData_ofExpr(x_246); x_352 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_352, 0, x_350); lean_ctor_set(x_352, 1, x_351); +x_353 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_354 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); lean_inc(x_12); -x_353 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_352, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_353; -goto block_277; +x_355 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_354, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_355; +goto block_279; } } } } 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; -lean_dec(x_246); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; +lean_dec(x_248); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_354 = l_Lean_MessageData_ofExpr(x_235); -x_355 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_356 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_354); -x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_356 = l_Lean_MessageData_ofExpr(x_236); +x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_358 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_358, 0, x_356); -lean_ctor_set(x_358, 1, x_357); -x_359 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_356); +x_359 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_360 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_360, 0, x_358); lean_ctor_set(x_360, 1, x_359); -x_361 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_361 = l_Lean_MessageData_ofExpr(x_246); x_362 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_362, 0, x_360); lean_ctor_set(x_362, 1, x_361); +x_363 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_364 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set(x_364, 1, x_363); lean_inc(x_12); -x_363 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_362, x_12, x_13, x_14, x_15, x_16, x_17, x_245); -x_247 = x_363; -goto block_277; +x_365 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_247); +x_249 = x_365; +goto block_279; } -block_277: +block_279: { -if (lean_obj_tag(x_247) == 0) +if (lean_obj_tag(x_249) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_250 = x_247; -} else { - lean_dec_ref(x_247); - x_250 = lean_box(0); -} -x_251 = lean_ctor_get(x_248, 0); +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); lean_inc(x_251); -x_252 = lean_ctor_get(x_248, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_253 = x_248; -} else { - lean_dec_ref(x_248); - x_253 = lean_box(0); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_252 = x_249; +} else { + lean_dec_ref(x_249); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_250, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_255 = x_250; +} else { + lean_dec_ref(x_250); + x_255 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_253)) { - x_254 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(0, 2, 0); } else { - x_254 = x_253; + x_256 = x_255; } -lean_ctor_set(x_254, 0, x_251); -lean_ctor_set(x_254, 1, x_6); -x_255 = lean_array_get_size(x_252); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_nat_dec_lt(x_256, x_255); -if (x_257 == 0) +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_6); +x_257 = lean_array_get_size(x_254); +x_258 = lean_unsigned_to_nat(0u); +x_259 = lean_nat_dec_lt(x_258, x_257); +if (x_259 == 0) { -lean_object* x_258; -lean_dec(x_255); +lean_object* x_260; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12588,25 +12593,25 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_258 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_260 = lean_alloc_ctor(0, 2, 0); } else { - x_258 = x_250; + x_260 = x_252; } -lean_ctor_set(x_258, 0, x_6); -lean_ctor_set(x_258, 1, x_249); -return x_258; +lean_ctor_set(x_260, 0, x_6); +lean_ctor_set(x_260, 1, x_251); +return x_260; } else { -uint8_t x_259; -x_259 = lean_nat_dec_le(x_255, x_255); -if (x_259 == 0) +uint8_t x_261; +x_261 = lean_nat_dec_le(x_257, x_257); +if (x_261 == 0) { -lean_object* x_260; -lean_dec(x_255); +lean_object* x_262; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12616,82 +12621,82 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_260 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_262 = lean_alloc_ctor(0, 2, 0); } else { - x_260 = x_250; + x_262 = x_252; } -lean_ctor_set(x_260, 0, x_6); -lean_ctor_set(x_260, 1, x_249); -return x_260; +lean_ctor_set(x_262, 0, x_6); +lean_ctor_set(x_262, 1, x_251); +return x_262; } else { -size_t x_261; size_t x_262; lean_object* x_263; -lean_dec(x_250); -lean_dec(x_6); -x_261 = 0; -x_262 = lean_usize_of_nat(x_255); -lean_dec(x_255); -x_263 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_252, x_261, x_262, x_254, x_12, x_13, x_14, x_15, x_16, x_17, x_249); +size_t x_263; size_t x_264; lean_object* x_265; lean_dec(x_252); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_266 = x_263; -} else { - lean_dec_ref(x_263); - x_266 = lean_box(0); -} -x_267 = lean_ctor_get(x_264, 1); +lean_dec(x_6); +x_263 = 0; +x_264 = lean_usize_of_nat(x_257); +lean_dec(x_257); +x_265 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__6___rarg(x_4, x_7, x_8, x_254, x_263, x_264, x_256, x_12, x_13, x_14, x_15, x_16, x_17, x_251); +lean_dec(x_254); +if (lean_obj_tag(x_265) == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_265, 1); lean_inc(x_267); -lean_dec(x_264); -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_268 = x_265; +} else { + lean_dec_ref(x_265); + x_268 = lean_box(0); +} +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +lean_dec(x_266); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(0, 2, 0); } else { - x_268 = x_266; + x_270 = x_268; } -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_265); -return x_268; +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_267); +return x_270; } else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_269 = lean_ctor_get(x_263, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_263, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_271 = x_263; +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_271 = lean_ctor_get(x_265, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_265, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_273 = x_265; } else { - lean_dec_ref(x_263); - x_271 = lean_box(0); + lean_dec_ref(x_265); + x_273 = lean_box(0); } -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_273)) { + x_274 = lean_alloc_ctor(1, 2, 0); } else { - x_272 = x_271; + x_274 = x_273; } -lean_ctor_set(x_272, 0, x_269); -lean_ctor_set(x_272, 1, x_270); -return x_272; +lean_ctor_set(x_274, 0, x_271); +lean_ctor_set(x_274, 1, x_272); +return x_274; } } } } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12702,33 +12707,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_273 = lean_ctor_get(x_247, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_247, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_275 = x_247; +x_275 = lean_ctor_get(x_249, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_249, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_277 = x_249; } else { - lean_dec_ref(x_247); - x_275 = lean_box(0); + lean_dec_ref(x_249); + x_277 = lean_box(0); } -if (lean_is_scalar(x_275)) { - x_276 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_276 = x_275; + x_278 = x_277; } -lean_ctor_set(x_276, 0, x_273); -lean_ctor_set(x_276, 1, x_274); -return x_276; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -lean_dec(x_235); +lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12744,32 +12749,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_364 = lean_ctor_get(x_243, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_243, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_243)) { - lean_ctor_release(x_243, 0); - lean_ctor_release(x_243, 1); - x_366 = x_243; +x_366 = lean_ctor_get(x_245, 0); +lean_inc(x_366); +x_367 = lean_ctor_get(x_245, 1); +lean_inc(x_367); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_368 = x_245; } else { - lean_dec_ref(x_243); - x_366 = lean_box(0); + lean_dec_ref(x_245); + x_368 = lean_box(0); } -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_368)) { + x_369 = lean_alloc_ctor(1, 2, 0); } else { - x_367 = x_366; + x_369 = x_368; } -lean_ctor_set(x_367, 0, x_364); -lean_ctor_set(x_367, 1, x_365); -return x_367; +lean_ctor_set(x_369, 0, x_366); +lean_ctor_set(x_369, 1, x_367); +return x_369; } } else { -lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; -lean_dec(x_235); +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12785,32 +12790,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_368 = lean_ctor_get(x_240, 0); -lean_inc(x_368); -x_369 = lean_ctor_get(x_240, 1); -lean_inc(x_369); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_370 = x_240; +x_370 = lean_ctor_get(x_242, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_242, 1); +lean_inc(x_371); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_372 = x_242; } else { - lean_dec_ref(x_240); - x_370 = lean_box(0); + lean_dec_ref(x_242); + x_372 = lean_box(0); } -if (lean_is_scalar(x_370)) { - x_371 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_372)) { + x_373 = lean_alloc_ctor(1, 2, 0); } else { - x_371 = x_370; + x_373 = x_372; } -lean_ctor_set(x_371, 0, x_368); -lean_ctor_set(x_371, 1, x_369); -return x_371; +lean_ctor_set(x_373, 0, x_370); +lean_ctor_set(x_373, 1, x_371); +return x_373; } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_235); +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12826,32 +12831,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_372 = lean_ctor_get(x_238, 0); -lean_inc(x_372); -x_373 = lean_ctor_get(x_238, 1); -lean_inc(x_373); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - lean_ctor_release(x_238, 1); - x_374 = x_238; +x_374 = lean_ctor_get(x_240, 0); +lean_inc(x_374); +x_375 = lean_ctor_get(x_240, 1); +lean_inc(x_375); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_376 = x_240; } else { - lean_dec_ref(x_238); - x_374 = lean_box(0); + lean_dec_ref(x_240); + x_376 = lean_box(0); } -if (lean_is_scalar(x_374)) { - x_375 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_376)) { + x_377 = lean_alloc_ctor(1, 2, 0); } else { - x_375 = x_374; + x_377 = x_376; } -lean_ctor_set(x_375, 0, x_372); -lean_ctor_set(x_375, 1, x_373); -return x_375; +lean_ctor_set(x_377, 0, x_374); +lean_ctor_set(x_377, 1, x_375); +return x_377; } } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -lean_dec(x_235); +lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12867,33 +12872,33 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_376 = lean_ctor_get(x_236, 0); -lean_inc(x_376); -x_377 = lean_ctor_get(x_236, 1); -lean_inc(x_377); -if (lean_is_exclusive(x_236)) { - lean_ctor_release(x_236, 0); - lean_ctor_release(x_236, 1); - x_378 = x_236; +x_378 = lean_ctor_get(x_237, 0); +lean_inc(x_378); +x_379 = lean_ctor_get(x_237, 1); +lean_inc(x_379); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_380 = x_237; } else { - lean_dec_ref(x_236); - x_378 = lean_box(0); + lean_dec_ref(x_237); + x_380 = lean_box(0); } -if (lean_is_scalar(x_378)) { - x_379 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_380)) { + x_381 = lean_alloc_ctor(1, 2, 0); } else { - x_379 = x_378; + x_381 = x_380; } -lean_ctor_set(x_379, 0, x_376); -lean_ctor_set(x_379, 1, x_377); -return x_379; +lean_ctor_set(x_381, 0, x_378); +lean_ctor_set(x_381, 1, x_379); +return x_381; } } } } else { -uint8_t x_389; +uint8_t x_391; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -12910,23 +12915,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_389 = !lean_is_exclusive(x_21); -if (x_389 == 0) +x_391 = !lean_is_exclusive(x_21); +if (x_391 == 0) { return x_21; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_21, 0); -x_391 = lean_ctor_get(x_21, 1); -lean_inc(x_391); -lean_inc(x_390); +lean_object* x_392; lean_object* x_393; lean_object* x_394; +x_392 = lean_ctor_get(x_21, 0); +x_393 = lean_ctor_get(x_21, 1); +lean_inc(x_393); +lean_inc(x_392); lean_dec(x_21); -x_392 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_392, 0, x_390); -lean_ctor_set(x_392, 1, x_391); -return x_392; +x_394 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_394, 0, x_392); +lean_ctor_set(x_394, 1, x_393); +return x_394; } } } @@ -12959,148 +12964,149 @@ lean_dec(x_24); x_25 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); if (lean_obj_tag(x_25) == 1) { -lean_object* x_227; -x_227 = lean_ctor_get(x_25, 0); -lean_inc(x_227); -lean_ctor_set(x_21, 0, x_227); +lean_object* x_228; +x_228 = lean_ctor_get(x_25, 0); +lean_inc(x_228); +lean_ctor_set(x_21, 0, x_228); x_26 = x_21; -goto block_226; +goto block_227; } else { -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_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_free_object(x_21); lean_inc(x_25); -x_228 = l_Lean_MessageData_ofExpr(x_25); -x_229 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_230 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_228); -x_231 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; -x_232 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_232, 0, x_230); -lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_232, x_14, x_15, x_16, x_17, x_23); -x_26 = x_233; -goto block_226; -} -block_226: +x_229 = l_Lean_MessageData_ofExpr(x_25); +x_230 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_231 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_229); +x_232 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_233 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_233, x_14, x_15, x_16, x_17, x_23); +x_26 = x_234; +goto block_227; +} +block_227: { if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_27; uint8_t x_28; lean_object* x_29; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); +x_28 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_28 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -if (lean_obj_tag(x_28) == 0) +x_29 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_28, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_25); -x_30 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_33 = l_Lean_Meta_whnfD(x_31, x_14, x_15, x_16, x_17, x_32); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_Meta_whnfD(x_32, x_14, x_15, x_16, x_17, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Expr_getAppFn(x_34); -if (lean_obj_tag(x_36) == 4) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Expr_getAppFn(x_35); +if (lean_obj_tag(x_37) == 4) { -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_ctor_get(x_36, 0); -lean_inc(x_110); -lean_dec(x_36); -x_111 = lean_st_ref_get(x_17, x_35); -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_111 = lean_ctor_get(x_37, 0); +lean_inc(x_111); +lean_dec(x_37); +x_112 = lean_st_ref_get(x_17, x_36); +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -lean_dec(x_111); -x_114 = lean_ctor_get(x_112, 0); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_110); -if (lean_obj_tag(x_115) == 0) +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_environment_find(x_115, x_111); +if (lean_obj_tag(x_116) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_object* x_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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_116 = l_Lean_MessageData_ofExpr(x_25); -x_117 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_116); -x_119 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_MessageData_ofExpr(x_34); -x_122 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_124 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); -lean_inc(x_12); -x_125 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_124, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_125; -goto block_109; +x_117 = l_Lean_MessageData_ofExpr(x_25); +x_118 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_119 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_121 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Lean_MessageData_ofExpr(x_35); +x_123 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_125 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); +lean_inc(x_12); +x_126 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_125, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_126; +goto block_110; } else { -lean_object* x_126; -x_126 = lean_ctor_get(x_115, 0); -lean_inc(x_126); -lean_dec(x_115); -switch (lean_obj_tag(x_126)) { +lean_object* x_127; +x_127 = lean_ctor_get(x_116, 0); +lean_inc(x_127); +lean_dec(x_116); +switch (lean_obj_tag(x_127)) { case 4: { -lean_object* x_127; lean_object* x_128; uint8_t x_144; lean_object* x_145; +lean_object* x_128; lean_object* x_129; uint8_t x_145; lean_object* x_146; lean_dec(x_1); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -lean_dec(x_126); -x_144 = lean_ctor_get_uint8(x_127, sizeof(void*)*1); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); lean_dec(x_127); -x_145 = lean_box(x_144); -if (lean_obj_tag(x_145) == 0) +x_145 = lean_ctor_get_uint8(x_128, sizeof(void*)*1); +lean_dec(x_128); +x_146 = lean_box(x_145); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_34); -x_146 = lean_box(0); +lean_object* x_147; lean_object* x_148; +lean_dec(x_35); +x_147 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -13108,357 +13114,357 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_147 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_146, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_147; -goto block_109; +x_148 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_148; +goto block_110; } else { -lean_object* x_148; -lean_dec(x_145); +lean_object* x_149; +lean_dec(x_146); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_148 = lean_box(0); -x_128 = x_148; -goto block_143; +x_149 = lean_box(0); +x_129 = x_149; +goto block_144; } -block_143: +block_144: { -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; uint8_t x_139; -lean_dec(x_128); -x_129 = l_Lean_MessageData_ofExpr(x_25); -x_130 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_131 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_133 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_MessageData_ofExpr(x_34); -x_135 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_137 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); +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; uint8_t x_140; +lean_dec(x_129); +x_130 = l_Lean_MessageData_ofExpr(x_25); +x_131 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_132 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_130); +x_133 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_134 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_MessageData_ofExpr(x_35); +x_136 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_138 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); lean_inc(x_12); -x_138 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_137, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_139 = !lean_is_exclusive(x_138); -if (x_139 == 0) +x_139 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_138, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) { -x_37 = x_138; -goto block_109; +x_38 = x_139; +goto block_110; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_138, 0); -x_141 = lean_ctor_get(x_138, 1); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_139, 0); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_138); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_37 = x_142; -goto block_109; +lean_dec(x_139); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_38 = x_143; +goto block_110; } } } case 5: { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_34); +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_35); lean_dec(x_11); lean_dec(x_3); -x_149 = lean_ctor_get(x_126, 0); -lean_inc(x_149); -lean_dec(x_126); -x_150 = lean_ctor_get(x_149, 1); +x_150 = lean_ctor_get(x_127, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 4); +lean_dec(x_127); +x_151 = lean_ctor_get(x_150, 1); lean_inc(x_151); -lean_dec(x_149); -x_152 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_153 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +x_152 = lean_ctor_get(x_150, 4); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_154 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_154 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_150, x_153, x_151, x_152, x_14, x_15, x_16, x_17, x_113); -lean_dec(x_150); -if (lean_obj_tag(x_154) == 0) +x_155 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_151, x_154, x_152, x_153, x_14, x_15, x_16, x_17, x_114); +lean_dec(x_151); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); -lean_dec(x_154); -x_157 = !lean_is_exclusive(x_155); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; -x_158 = lean_ctor_get(x_155, 0); -x_159 = lean_ctor_get(x_155, 1); -x_160 = l_Lean_Expr_fvarId_x21(x_25); -x_161 = 1; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_158 = !lean_is_exclusive(x_156); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; +x_159 = lean_ctor_get(x_156, 0); +x_160 = lean_ctor_get(x_156, 1); +x_161 = l_Lean_Expr_fvarId_x21(x_25); +x_162 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_162 = l_Lean_Meta_Cases_cases(x_10, x_160, x_158, x_161, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_162) == 0) +x_163 = l_Lean_Meta_Cases_cases(x_10, x_161, x_159, x_162, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_163) == 0) { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_162); -if (x_163 == 0) +uint8_t x_164; +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) { -lean_object* x_164; -x_164 = lean_ctor_get(x_162, 0); -lean_ctor_set(x_155, 1, x_164); -lean_ctor_set(x_155, 0, x_159); -lean_ctor_set(x_162, 0, x_155); -x_37 = x_162; -goto block_109; +lean_object* x_165; +x_165 = lean_ctor_get(x_163, 0); +lean_ctor_set(x_156, 1, x_165); +lean_ctor_set(x_156, 0, x_160); +lean_ctor_set(x_163, 0, x_156); +x_38 = x_163; +goto block_110; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_162, 0); -x_166 = lean_ctor_get(x_162, 1); +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_163, 0); +x_167 = lean_ctor_get(x_163, 1); +lean_inc(x_167); lean_inc(x_166); -lean_inc(x_165); -lean_dec(x_162); -lean_ctor_set(x_155, 1, x_165); -lean_ctor_set(x_155, 0, x_159); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_155); -lean_ctor_set(x_167, 1, x_166); -x_37 = x_167; -goto block_109; +lean_dec(x_163); +lean_ctor_set(x_156, 1, x_166); +lean_ctor_set(x_156, 0, x_160); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_156); +lean_ctor_set(x_168, 1, x_167); +x_38 = x_168; +goto block_110; } } else { -uint8_t x_168; -lean_free_object(x_155); -lean_dec(x_159); -x_168 = !lean_is_exclusive(x_162); -if (x_168 == 0) +uint8_t x_169; +lean_free_object(x_156); +lean_dec(x_160); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -x_37 = x_162; -goto block_109; +x_38 = x_163; +goto block_110; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_162, 0); -x_170 = lean_ctor_get(x_162, 1); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_163, 0); +x_171 = lean_ctor_get(x_163, 1); +lean_inc(x_171); lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_162); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_37 = x_171; -goto block_109; +lean_dec(x_163); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_38 = x_172; +goto block_110; } } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; -x_172 = lean_ctor_get(x_155, 0); -x_173 = lean_ctor_get(x_155, 1); +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; +x_173 = lean_ctor_get(x_156, 0); +x_174 = lean_ctor_get(x_156, 1); +lean_inc(x_174); lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_155); -x_174 = l_Lean_Expr_fvarId_x21(x_25); -x_175 = 1; +lean_dec(x_156); +x_175 = l_Lean_Expr_fvarId_x21(x_25); +x_176 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_176 = l_Lean_Meta_Cases_cases(x_10, x_174, x_172, x_175, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_176) == 0) +x_177 = l_Lean_Meta_Cases_cases(x_10, x_175, x_173, x_176, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_177) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_179 = x_176; +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_180 = x_177; } else { - lean_dec_ref(x_176); - x_179 = lean_box(0); + lean_dec_ref(x_177); + x_180 = lean_box(0); } -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_173); -lean_ctor_set(x_180, 1, x_177); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(0, 2, 0); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_174); +lean_ctor_set(x_181, 1, x_178); +if (lean_is_scalar(x_180)) { + x_182 = lean_alloc_ctor(0, 2, 0); } else { - x_181 = x_179; + x_182 = x_180; } -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_178); -x_37 = x_181; -goto block_109; +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +x_38 = x_182; +goto block_110; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_173); -x_182 = lean_ctor_get(x_176, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_176, 1); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_174); +x_183 = lean_ctor_get(x_177, 0); lean_inc(x_183); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_184 = x_176; +x_184 = lean_ctor_get(x_177, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_185 = x_177; } else { - lean_dec_ref(x_176); - x_184 = lean_box(0); + lean_dec_ref(x_177); + x_185 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_186 = x_185; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -x_37 = x_185; -goto block_109; +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_184); +x_38 = x_186; +goto block_110; } } } else { -uint8_t x_186; +uint8_t x_187; lean_dec(x_25); lean_dec(x_10); -x_186 = !lean_is_exclusive(x_154); -if (x_186 == 0) +x_187 = !lean_is_exclusive(x_155); +if (x_187 == 0) { -x_37 = x_154; -goto block_109; +x_38 = x_155; +goto block_110; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_154, 0); -x_188 = lean_ctor_get(x_154, 1); +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_155, 0); +x_189 = lean_ctor_get(x_155, 1); +lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_154); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -x_37 = x_189; -goto block_109; +lean_dec(x_155); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_38 = x_190; +goto block_110; } } } default: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_126); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_127); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_190 = l_Lean_MessageData_ofExpr(x_25); -x_191 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_192 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_190); -x_193 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_MessageData_ofExpr(x_34); -x_196 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -x_197 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_198 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); +x_191 = l_Lean_MessageData_ofExpr(x_25); +x_192 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_MessageData_ofExpr(x_35); +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 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_199 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); lean_inc(x_12); -x_199 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_199; -goto block_109; +x_200 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_199, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_200; +goto block_110; } } } } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_36); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_37); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_200 = l_Lean_MessageData_ofExpr(x_25); -x_201 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_202 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_200); -x_203 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_204 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -x_205 = l_Lean_MessageData_ofExpr(x_34); -x_206 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_208 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); +x_201 = l_Lean_MessageData_ofExpr(x_25); +x_202 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_203 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_205 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_MessageData_ofExpr(x_35); +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +x_208 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_209 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); lean_inc(x_12); -x_209 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_208, x_12, x_13, x_14, x_15, x_16, x_17, x_35); -x_37 = x_209; -goto block_109; +x_210 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_209, x_12, x_13, x_14, x_15, x_16, x_17, x_36); +x_38 = x_210; +goto block_110; } -block_109: +block_110: { -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_38) == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_38, 0); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_37, 1); -x_42 = lean_ctor_get(x_39, 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_38, 1); +x_43 = lean_ctor_get(x_40, 1); lean_inc(x_6); -lean_ctor_set(x_39, 1, x_6); -x_43 = lean_array_get_size(x_42); -x_44 = lean_unsigned_to_nat(0u); -x_45 = lean_nat_dec_lt(x_44, x_43); -if (x_45 == 0) +lean_ctor_set(x_40, 1, x_6); +x_44 = lean_array_get_size(x_43); +x_45 = lean_unsigned_to_nat(0u); +x_46 = lean_nat_dec_lt(x_45, x_44); +if (x_46 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13468,18 +13474,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_43, x_43); -if (x_46 == 0) +uint8_t x_47; +x_47 = lean_nat_dec_le(x_44, x_44); +if (x_47 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13489,70 +13495,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_47; size_t x_48; lean_object* x_49; -lean_free_object(x_37); +size_t x_48; size_t x_49; lean_object* x_50; +lean_free_object(x_38); lean_dec(x_6); -x_47 = 0; -x_48 = lean_usize_of_nat(x_43); +x_48 = 0; +x_49 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_43, x_48, x_49, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_42); lean_dec(x_43); -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_42, x_47, x_48, x_39, x_12, x_13, x_14, x_15, x_16, x_17, x_41); -lean_dec(x_42); -if (lean_obj_tag(x_49) == 0) +if (lean_obj_tag(x_50) == 0) { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_52); -return x_49; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_53); +return x_50; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_49, 0); -x_54 = lean_ctor_get(x_49, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_49); -x_55 = lean_ctor_get(x_53, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +x_55 = lean_ctor_get(x_50, 1); lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_inc(x_54); +lean_dec(x_50); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_58; +x_58 = !lean_is_exclusive(x_50); +if (x_58 == 0) { -return x_49; +return x_50; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_50, 0); +x_60 = lean_ctor_get(x_50, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_50); +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; } } } @@ -13560,25 +13566,25 @@ return x_60; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_61 = lean_ctor_get(x_37, 1); -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_62 = lean_ctor_get(x_38, 1); +x_63 = lean_ctor_get(x_40, 0); +x_64 = lean_ctor_get(x_40, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); +lean_dec(x_40); lean_inc(x_6); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_6); -x_65 = lean_array_get_size(x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_6); +x_66 = lean_array_get_size(x_64); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +if (x_68 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13588,18 +13594,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_68; -x_68 = lean_nat_dec_le(x_65, x_65); -if (x_68 == 0) +uint8_t x_69; +x_69 = lean_nat_dec_le(x_66, x_66); +if (x_69 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13609,69 +13615,69 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_69; size_t x_70; lean_object* x_71; -lean_free_object(x_37); +size_t x_70; size_t x_71; lean_object* x_72; +lean_free_object(x_38); lean_dec(x_6); -x_69 = 0; -x_70 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_71 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_63, x_69, x_70, x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -lean_dec(x_63); -if (lean_obj_tag(x_71) == 0) +x_70 = 0; +x_71 = lean_usize_of_nat(x_66); +lean_dec(x_66); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_64, x_70, x_71, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_64); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); 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_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_72); + x_75 = lean_box(0); } -x_75 = lean_ctor_get(x_72, 1); -lean_inc(x_75); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_76 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_76 = x_74; + x_77 = x_75; } -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_73); -return x_76; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_71, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_71, 1); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_72, 0); lean_inc(x_78); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_79 = x_71; +x_79 = lean_ctor_get(x_72, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_80 = x_72; } else { - lean_dec_ref(x_71); - x_79 = lean_box(0); + lean_dec_ref(x_72); + x_80 = lean_box(0); } -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_80 = x_79; + x_81 = x_80; } -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_78); -return x_80; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } @@ -13679,41 +13685,41 @@ 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; uint8_t x_89; -x_81 = lean_ctor_get(x_37, 0); -x_82 = lean_ctor_get(x_37, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_37); -x_83 = lean_ctor_get(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; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_82 = lean_ctor_get(x_38, 0); +x_83 = lean_ctor_get(x_38, 1); lean_inc(x_83); -x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_38); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_85 = x_81; +x_85 = lean_ctor_get(x_82, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_81); - x_85 = lean_box(0); + lean_dec_ref(x_82); + x_86 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_86 = x_85; + x_87 = x_86; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_6); -x_87 = lean_array_get_size(x_84); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_nat_dec_lt(x_88, x_87); -if (x_89 == 0) +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_6); +x_88 = lean_array_get_size(x_85); +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_nat_dec_lt(x_89, x_88); +if (x_90 == 0) { -lean_object* x_90; +lean_object* x_91; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13723,21 +13729,21 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_82); -return x_90; +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_6); +lean_ctor_set(x_91, 1, x_83); +return x_91; } else { -uint8_t x_91; -x_91 = lean_nat_dec_le(x_87, x_87); -if (x_91 == 0) +uint8_t x_92; +x_92 = lean_nat_dec_le(x_88, x_88); +if (x_92 == 0) { -lean_object* x_92; +lean_object* x_93; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13747,70 +13753,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_6); -lean_ctor_set(x_92, 1, x_82); -return x_92; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_6); +lean_ctor_set(x_93, 1, x_83); +return x_93; } else { -size_t x_93; size_t x_94; lean_object* x_95; +size_t x_94; size_t x_95; lean_object* x_96; lean_dec(x_6); -x_93 = 0; -x_94 = lean_usize_of_nat(x_87); -lean_dec(x_87); -x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_84, x_93, x_94, x_86, x_12, x_13, x_14, x_15, x_16, x_17, x_82); -lean_dec(x_84); -if (lean_obj_tag(x_95) == 0) +x_94 = 0; +x_95 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_96 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_85, x_94, x_95, x_87, x_12, x_13, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_85); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_ctor_get(x_96, 0); lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_99 = x_96; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_96); + x_99 = lean_box(0); } -x_99 = lean_ctor_get(x_96, 1); -lean_inc(x_99); -lean_dec(x_96); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_101 = x_99; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_101 = lean_ctor_get(x_95, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_95, 1); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_96, 0); lean_inc(x_102); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_103 = x_95; +x_103 = lean_ctor_get(x_96, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_104 = x_96; } else { - lean_dec_ref(x_95); - x_103 = lean_box(0); + lean_dec_ref(x_96); + x_104 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_105 = x_104; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } } @@ -13818,7 +13824,7 @@ return x_104; } else { -uint8_t x_105; +uint8_t x_106; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -13829,30 +13835,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_37); -if (x_105 == 0) +x_106 = !lean_is_exclusive(x_38); +if (x_106 == 0) { -return x_37; +return x_38; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_37, 0); -x_107 = lean_ctor_get(x_37, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_38, 0); +x_108 = lean_ctor_get(x_38, 1); +lean_inc(x_108); lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_37); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_dec(x_38); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; } } } } else { -uint8_t x_210; +uint8_t x_211; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -13869,29 +13875,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_210 = !lean_is_exclusive(x_33); -if (x_210 == 0) +x_211 = !lean_is_exclusive(x_34); +if (x_211 == 0) { -return x_33; +return x_34; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_33, 0); -x_212 = lean_ctor_get(x_33, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_34, 0); +x_213 = lean_ctor_get(x_34, 1); +lean_inc(x_213); lean_inc(x_212); -lean_inc(x_211); -lean_dec(x_33); -x_213 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -return x_213; +lean_dec(x_34); +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +return x_214; } } } else { -uint8_t x_214; +uint8_t x_215; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -13908,29 +13914,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_214 = !lean_is_exclusive(x_30); -if (x_214 == 0) +x_215 = !lean_is_exclusive(x_31); +if (x_215 == 0) { -return x_30; +return x_31; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_215 = lean_ctor_get(x_30, 0); -x_216 = lean_ctor_get(x_30, 1); +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_ctor_get(x_31, 0); +x_217 = lean_ctor_get(x_31, 1); +lean_inc(x_217); lean_inc(x_216); -lean_inc(x_215); -lean_dec(x_30); -x_217 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -return x_217; +lean_dec(x_31); +x_218 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +return x_218; } } } else { -uint8_t x_218; +uint8_t x_219; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -13947,29 +13953,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_218 = !lean_is_exclusive(x_28); -if (x_218 == 0) +x_219 = !lean_is_exclusive(x_29); +if (x_219 == 0) { -return x_28; +return x_29; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_219 = lean_ctor_get(x_28, 0); -x_220 = lean_ctor_get(x_28, 1); +lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_220 = lean_ctor_get(x_29, 0); +x_221 = lean_ctor_get(x_29, 1); +lean_inc(x_221); lean_inc(x_220); -lean_inc(x_219); -lean_dec(x_28); -x_221 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -return x_221; +lean_dec(x_29); +x_222 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } } else { -uint8_t x_222; +uint8_t x_223; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -13986,179 +13992,180 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_222 = !lean_is_exclusive(x_26); -if (x_222 == 0) +x_223 = !lean_is_exclusive(x_26); +if (x_223 == 0) { return x_26; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_26, 0); -x_224 = lean_ctor_get(x_26, 1); +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_26, 0); +x_225 = lean_ctor_get(x_26, 1); +lean_inc(x_225); lean_inc(x_224); -lean_inc(x_223); lean_dec(x_26); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +x_226 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_226, 0, x_224); +lean_ctor_set(x_226, 1, x_225); +return x_226; } } } } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_234 = lean_ctor_get(x_21, 1); -lean_inc(x_234); +lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_235 = lean_ctor_get(x_21, 1); +lean_inc(x_235); lean_dec(x_21); -x_235 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); -if (lean_obj_tag(x_235) == 1) +x_236 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); +if (lean_obj_tag(x_236) == 1) { -lean_object* x_381; lean_object* x_382; -x_381 = lean_ctor_get(x_235, 0); -lean_inc(x_381); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_381); -lean_ctor_set(x_382, 1, x_234); -x_236 = x_382; -goto block_380; +lean_object* x_383; lean_object* x_384; +x_383 = lean_ctor_get(x_236, 0); +lean_inc(x_383); +x_384 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_235); +x_237 = x_384; +goto block_382; } else { -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_inc(x_235); -x_383 = l_Lean_MessageData_ofExpr(x_235); -x_384 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_385 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_383); -x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +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_inc(x_236); +x_385 = l_Lean_MessageData_ofExpr(x_236); +x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_387 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_387, 0, x_385); -lean_ctor_set(x_387, 1, x_386); -x_388 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_387, x_14, x_15, x_16, x_17, x_234); -x_236 = x_388; -goto block_380; -} -block_380: -{ -if (lean_obj_tag(x_236) == 0) -{ -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_236, 1); -lean_inc(x_237); -lean_dec(x_236); +lean_ctor_set(x_387, 0, x_386); +lean_ctor_set(x_387, 1, x_385); +x_388 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_389 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_389, 0, x_387); +lean_ctor_set(x_389, 1, x_388); +x_390 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_389, x_14, x_15, x_16, x_17, x_235); +x_237 = x_390; +goto block_382; +} +block_382: +{ +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; uint8_t x_239; lean_object* x_240; +x_238 = lean_ctor_get(x_237, 1); +lean_inc(x_238); +lean_dec(x_237); +x_239 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_238 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_237); -if (lean_obj_tag(x_238) == 0) -{ -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_238, 1); -lean_inc(x_239); -lean_dec(x_238); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_235); -x_240 = lean_infer_type(x_235, x_14, x_15, x_16, x_17, x_239); +x_240 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_239, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_238); if (lean_obj_tag(x_240) == 0) { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_240, 0); +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_240, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); lean_dec(x_240); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_243 = l_Lean_Meta_whnfD(x_241, x_14, x_15, x_16, x_17, x_242); -if (lean_obj_tag(x_243) == 0) +lean_inc(x_236); +x_242 = lean_infer_type(x_236, x_14, x_15, x_16, x_17, x_241); +if (lean_obj_tag(x_242) == 0) { -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_244 = lean_ctor_get(x_243, 0); +lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); -x_245 = lean_ctor_get(x_243, 1); -lean_inc(x_245); -lean_dec(x_243); -x_246 = l_Lean_Expr_getAppFn(x_244); -if (lean_obj_tag(x_246) == 4) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_278 = lean_ctor_get(x_246, 0); -lean_inc(x_278); -lean_dec(x_246); -x_279 = lean_st_ref_get(x_17, x_245); -x_280 = lean_ctor_get(x_279, 0); +lean_dec(x_242); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_245 = l_Lean_Meta_whnfD(x_243, x_14, x_15, x_16, x_17, x_244); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec(x_245); +x_248 = l_Lean_Expr_getAppFn(x_246); +if (lean_obj_tag(x_248) == 4) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_280 = lean_ctor_get(x_248, 0); lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_282 = lean_ctor_get(x_280, 0); +lean_dec(x_248); +x_281 = lean_st_ref_get(x_17, x_247); +x_282 = lean_ctor_get(x_281, 0); lean_inc(x_282); -lean_dec(x_280); -x_283 = lean_environment_find(x_282, x_278); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_ctor_get(x_282, 0); +lean_inc(x_284); +lean_dec(x_282); +x_285 = lean_environment_find(x_284, x_280); +if (lean_obj_tag(x_285) == 0) +{ +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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_284 = l_Lean_MessageData_ofExpr(x_235); -x_285 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_286 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_284); -x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_286 = l_Lean_MessageData_ofExpr(x_236); +x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_288 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_288, 0, x_286); -lean_ctor_set(x_288, 1, x_287); -x_289 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_286); +x_289 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_290 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_290, 0, x_288); lean_ctor_set(x_290, 1, x_289); -x_291 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_291 = l_Lean_MessageData_ofExpr(x_246); x_292 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_292, 0, x_290); lean_ctor_set(x_292, 1, x_291); +x_293 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_294 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); lean_inc(x_12); -x_293 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_293; -goto block_277; +x_295 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_294, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_295; +goto block_279; } else { -lean_object* x_294; -x_294 = lean_ctor_get(x_283, 0); -lean_inc(x_294); -lean_dec(x_283); -switch (lean_obj_tag(x_294)) { +lean_object* x_296; +x_296 = lean_ctor_get(x_285, 0); +lean_inc(x_296); +lean_dec(x_285); +switch (lean_obj_tag(x_296)) { case 4: { -lean_object* x_295; lean_object* x_296; uint8_t x_312; lean_object* x_313; +lean_object* x_297; lean_object* x_298; uint8_t x_314; lean_object* x_315; lean_dec(x_1); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -lean_dec(x_294); -x_312 = lean_ctor_get_uint8(x_295, sizeof(void*)*1); -lean_dec(x_295); -x_313 = lean_box(x_312); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; lean_object* x_315; -lean_dec(x_244); -x_314 = lean_box(0); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +lean_dec(x_296); +x_314 = lean_ctor_get_uint8(x_297, sizeof(void*)*1); +lean_dec(x_297); +x_315 = lean_box(x_314); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; +lean_dec(x_246); +x_316 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -14166,316 +14173,316 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_315 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_235, x_10, x_3, x_11, x_4, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_315; -goto block_277; +x_317 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_236, x_10, x_3, x_11, x_4, x_316, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_317; +goto block_279; } else { -lean_object* x_316; -lean_dec(x_313); +lean_object* x_318; +lean_dec(x_315); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_316 = lean_box(0); -x_296 = x_316; -goto block_311; +x_318 = lean_box(0); +x_298 = x_318; +goto block_313; } -block_311: +block_313: { -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_296); -x_297 = l_Lean_MessageData_ofExpr(x_235); -x_298 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_299 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +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_dec(x_298); +x_299 = l_Lean_MessageData_ofExpr(x_236); +x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_301 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_299); +x_302 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_303 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_303, 0, x_301); lean_ctor_set(x_303, 1, x_302); -x_304 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_304 = l_Lean_MessageData_ofExpr(x_246); x_305 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_305, 0, x_303); lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_307 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_307, 0, x_305); +lean_ctor_set(x_307, 1, x_306); lean_inc(x_12); -x_306 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_305, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_307 = lean_ctor_get(x_306, 0); -lean_inc(x_307); -x_308 = lean_ctor_get(x_306, 1); -lean_inc(x_308); -if (lean_is_exclusive(x_306)) { - lean_ctor_release(x_306, 0); - lean_ctor_release(x_306, 1); - x_309 = x_306; +x_308 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_307, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_311 = x_308; } else { - lean_dec_ref(x_306); - x_309 = lean_box(0); + lean_dec_ref(x_308); + x_311 = lean_box(0); } -if (lean_is_scalar(x_309)) { - x_310 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_311)) { + x_312 = lean_alloc_ctor(1, 2, 0); } else { - x_310 = x_309; + x_312 = x_311; } -lean_ctor_set(x_310, 0, x_307); -lean_ctor_set(x_310, 1, x_308); -x_247 = x_310; -goto block_277; +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_310); +x_249 = x_312; +goto block_279; } } case 5: { -lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_244); +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_246); lean_dec(x_11); lean_dec(x_3); -x_317 = lean_ctor_get(x_294, 0); -lean_inc(x_317); -lean_dec(x_294); -x_318 = lean_ctor_get(x_317, 1); -lean_inc(x_318); -x_319 = lean_ctor_get(x_317, 4); +x_319 = lean_ctor_get(x_296, 0); lean_inc(x_319); -lean_dec(x_317); -x_320 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_321 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +lean_dec(x_296); +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +x_321 = lean_ctor_get(x_319, 4); +lean_inc(x_321); +lean_dec(x_319); +x_322 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_323 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_322 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_318, x_321, x_319, x_320, x_14, x_15, x_16, x_17, x_281); -lean_dec(x_318); -if (lean_obj_tag(x_322) == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; lean_object* x_330; -x_323 = lean_ctor_get(x_322, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_322, 1); -lean_inc(x_324); -lean_dec(x_322); -x_325 = lean_ctor_get(x_323, 0); +x_324 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_320, x_323, x_321, x_322, x_14, x_15, x_16, x_17, x_283); +lean_dec(x_320); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; lean_object* x_332; +x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); -x_326 = lean_ctor_get(x_323, 1); +x_326 = lean_ctor_get(x_324, 1); lean_inc(x_326); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_327 = x_323; -} else { - lean_dec_ref(x_323); - x_327 = lean_box(0); -} -x_328 = l_Lean_Expr_fvarId_x21(x_235); -x_329 = 1; +lean_dec(x_324); +x_327 = lean_ctor_get(x_325, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_329 = x_325; +} else { + lean_dec_ref(x_325); + x_329 = lean_box(0); +} +x_330 = l_Lean_Expr_fvarId_x21(x_236); +x_331 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_330 = l_Lean_Meta_Cases_cases(x_10, x_328, x_325, x_329, x_14, x_15, x_16, x_17, x_324); -if (lean_obj_tag(x_330) == 0) +x_332 = l_Lean_Meta_Cases_cases(x_10, x_330, x_327, x_331, x_14, x_15, x_16, x_17, x_326); +if (lean_obj_tag(x_332) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_330, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_333 = x_330; +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_332, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_335 = x_332; } else { - lean_dec_ref(x_330); - x_333 = lean_box(0); + lean_dec_ref(x_332); + x_335 = lean_box(0); } -if (lean_is_scalar(x_327)) { - x_334 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_329)) { + x_336 = lean_alloc_ctor(0, 2, 0); } else { - x_334 = x_327; + x_336 = x_329; } -lean_ctor_set(x_334, 0, x_326); -lean_ctor_set(x_334, 1, x_331); -if (lean_is_scalar(x_333)) { - x_335 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_328); +lean_ctor_set(x_336, 1, x_333); +if (lean_is_scalar(x_335)) { + x_337 = lean_alloc_ctor(0, 2, 0); } else { - x_335 = x_333; + x_337 = x_335; } -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_332); -x_247 = x_335; -goto block_277; +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_334); +x_249 = x_337; +goto block_279; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_327); -lean_dec(x_326); -x_336 = lean_ctor_get(x_330, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_330, 1); -lean_inc(x_337); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_338 = x_330; +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_329); +lean_dec(x_328); +x_338 = lean_ctor_get(x_332, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_332, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_340 = x_332; } else { - lean_dec_ref(x_330); - x_338 = lean_box(0); + lean_dec_ref(x_332); + x_340 = lean_box(0); } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); } else { - x_339 = x_338; + x_341 = x_340; } -lean_ctor_set(x_339, 0, x_336); -lean_ctor_set(x_339, 1, x_337); -x_247 = x_339; -goto block_277; +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +x_249 = x_341; +goto block_279; } } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_235); +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_236); lean_dec(x_10); -x_340 = lean_ctor_get(x_322, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_322, 1); -lean_inc(x_341); -if (lean_is_exclusive(x_322)) { - lean_ctor_release(x_322, 0); - lean_ctor_release(x_322, 1); - x_342 = x_322; +x_342 = lean_ctor_get(x_324, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_324, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_344 = x_324; } else { - lean_dec_ref(x_322); - x_342 = lean_box(0); + lean_dec_ref(x_324); + x_344 = lean_box(0); } -if (lean_is_scalar(x_342)) { - x_343 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); } else { - x_343 = x_342; + x_345 = x_344; } -lean_ctor_set(x_343, 0, x_340); -lean_ctor_set(x_343, 1, x_341); -x_247 = x_343; -goto block_277; +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +x_249 = x_345; +goto block_279; } } default: { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -lean_dec(x_294); +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_296); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_344 = l_Lean_MessageData_ofExpr(x_235); -x_345 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_346 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_344); -x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_346 = l_Lean_MessageData_ofExpr(x_236); +x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_348 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_348, 0, x_346); -lean_ctor_set(x_348, 1, x_347); -x_349 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_348, 0, x_347); +lean_ctor_set(x_348, 1, x_346); +x_349 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_350 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_350, 0, x_348); lean_ctor_set(x_350, 1, x_349); -x_351 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_351 = l_Lean_MessageData_ofExpr(x_246); x_352 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_352, 0, x_350); lean_ctor_set(x_352, 1, x_351); +x_353 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_354 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); lean_inc(x_12); -x_353 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_352, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_353; -goto block_277; +x_355 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_354, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_355; +goto block_279; } } } } 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; -lean_dec(x_246); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; +lean_dec(x_248); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_354 = l_Lean_MessageData_ofExpr(x_235); -x_355 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_356 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_354); -x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_356 = l_Lean_MessageData_ofExpr(x_236); +x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_358 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_358, 0, x_356); -lean_ctor_set(x_358, 1, x_357); -x_359 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_356); +x_359 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_360 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_360, 0, x_358); lean_ctor_set(x_360, 1, x_359); -x_361 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_361 = l_Lean_MessageData_ofExpr(x_246); x_362 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_362, 0, x_360); lean_ctor_set(x_362, 1, x_361); +x_363 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_364 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set(x_364, 1, x_363); lean_inc(x_12); -x_363 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_362, x_12, x_13, x_14, x_15, x_16, x_17, x_245); -x_247 = x_363; -goto block_277; +x_365 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_247); +x_249 = x_365; +goto block_279; } -block_277: +block_279: { -if (lean_obj_tag(x_247) == 0) +if (lean_obj_tag(x_249) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_250 = x_247; -} else { - lean_dec_ref(x_247); - x_250 = lean_box(0); -} -x_251 = lean_ctor_get(x_248, 0); +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); lean_inc(x_251); -x_252 = lean_ctor_get(x_248, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_253 = x_248; -} else { - lean_dec_ref(x_248); - x_253 = lean_box(0); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_252 = x_249; +} else { + lean_dec_ref(x_249); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_250, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_255 = x_250; +} else { + lean_dec_ref(x_250); + x_255 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_253)) { - x_254 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(0, 2, 0); } else { - x_254 = x_253; + x_256 = x_255; } -lean_ctor_set(x_254, 0, x_251); -lean_ctor_set(x_254, 1, x_6); -x_255 = lean_array_get_size(x_252); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_nat_dec_lt(x_256, x_255); -if (x_257 == 0) +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_6); +x_257 = lean_array_get_size(x_254); +x_258 = lean_unsigned_to_nat(0u); +x_259 = lean_nat_dec_lt(x_258, x_257); +if (x_259 == 0) { -lean_object* x_258; -lean_dec(x_255); +lean_object* x_260; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14485,25 +14492,25 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_258 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_260 = lean_alloc_ctor(0, 2, 0); } else { - x_258 = x_250; + x_260 = x_252; } -lean_ctor_set(x_258, 0, x_6); -lean_ctor_set(x_258, 1, x_249); -return x_258; +lean_ctor_set(x_260, 0, x_6); +lean_ctor_set(x_260, 1, x_251); +return x_260; } else { -uint8_t x_259; -x_259 = lean_nat_dec_le(x_255, x_255); -if (x_259 == 0) +uint8_t x_261; +x_261 = lean_nat_dec_le(x_257, x_257); +if (x_261 == 0) { -lean_object* x_260; -lean_dec(x_255); +lean_object* x_262; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14513,82 +14520,82 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_260 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_262 = lean_alloc_ctor(0, 2, 0); } else { - x_260 = x_250; + x_262 = x_252; } -lean_ctor_set(x_260, 0, x_6); -lean_ctor_set(x_260, 1, x_249); -return x_260; +lean_ctor_set(x_262, 0, x_6); +lean_ctor_set(x_262, 1, x_251); +return x_262; } else { -size_t x_261; size_t x_262; lean_object* x_263; -lean_dec(x_250); -lean_dec(x_6); -x_261 = 0; -x_262 = lean_usize_of_nat(x_255); -lean_dec(x_255); -x_263 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_252, x_261, x_262, x_254, x_12, x_13, x_14, x_15, x_16, x_17, x_249); +size_t x_263; size_t x_264; lean_object* x_265; lean_dec(x_252); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_266 = x_263; -} else { - lean_dec_ref(x_263); - x_266 = lean_box(0); -} -x_267 = lean_ctor_get(x_264, 1); +lean_dec(x_6); +x_263 = 0; +x_264 = lean_usize_of_nat(x_257); +lean_dec(x_257); +x_265 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__7___rarg(x_4, x_7, x_8, x_254, x_263, x_264, x_256, x_12, x_13, x_14, x_15, x_16, x_17, x_251); +lean_dec(x_254); +if (lean_obj_tag(x_265) == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_265, 1); lean_inc(x_267); -lean_dec(x_264); -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_268 = x_265; } else { - x_268 = x_266; + lean_dec_ref(x_265); + x_268 = lean_box(0); } -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_265); -return x_268; +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +lean_dec(x_266); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(0, 2, 0); +} else { + x_270 = x_268; +} +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_267); +return x_270; } else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_269 = lean_ctor_get(x_263, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_263, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_271 = x_263; +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_271 = lean_ctor_get(x_265, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_265, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_273 = x_265; } else { - lean_dec_ref(x_263); - x_271 = lean_box(0); + lean_dec_ref(x_265); + x_273 = lean_box(0); } -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_273)) { + x_274 = lean_alloc_ctor(1, 2, 0); } else { - x_272 = x_271; + x_274 = x_273; } -lean_ctor_set(x_272, 0, x_269); -lean_ctor_set(x_272, 1, x_270); -return x_272; +lean_ctor_set(x_274, 0, x_271); +lean_ctor_set(x_274, 1, x_272); +return x_274; } } } } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14599,33 +14606,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_273 = lean_ctor_get(x_247, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_247, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_275 = x_247; +x_275 = lean_ctor_get(x_249, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_249, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_277 = x_249; } else { - lean_dec_ref(x_247); - x_275 = lean_box(0); + lean_dec_ref(x_249); + x_277 = lean_box(0); } -if (lean_is_scalar(x_275)) { - x_276 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_276 = x_275; + x_278 = x_277; } -lean_ctor_set(x_276, 0, x_273); -lean_ctor_set(x_276, 1, x_274); -return x_276; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -lean_dec(x_235); +lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14641,32 +14648,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_364 = lean_ctor_get(x_243, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_243, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_243)) { - lean_ctor_release(x_243, 0); - lean_ctor_release(x_243, 1); - x_366 = x_243; +x_366 = lean_ctor_get(x_245, 0); +lean_inc(x_366); +x_367 = lean_ctor_get(x_245, 1); +lean_inc(x_367); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_368 = x_245; } else { - lean_dec_ref(x_243); - x_366 = lean_box(0); + lean_dec_ref(x_245); + x_368 = lean_box(0); } -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_368)) { + x_369 = lean_alloc_ctor(1, 2, 0); } else { - x_367 = x_366; + x_369 = x_368; } -lean_ctor_set(x_367, 0, x_364); -lean_ctor_set(x_367, 1, x_365); -return x_367; +lean_ctor_set(x_369, 0, x_366); +lean_ctor_set(x_369, 1, x_367); +return x_369; } } else { -lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; -lean_dec(x_235); +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14682,32 +14689,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_368 = lean_ctor_get(x_240, 0); -lean_inc(x_368); -x_369 = lean_ctor_get(x_240, 1); -lean_inc(x_369); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_370 = x_240; +x_370 = lean_ctor_get(x_242, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_242, 1); +lean_inc(x_371); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_372 = x_242; } else { - lean_dec_ref(x_240); - x_370 = lean_box(0); + lean_dec_ref(x_242); + x_372 = lean_box(0); } -if (lean_is_scalar(x_370)) { - x_371 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_372)) { + x_373 = lean_alloc_ctor(1, 2, 0); } else { - x_371 = x_370; + x_373 = x_372; } -lean_ctor_set(x_371, 0, x_368); -lean_ctor_set(x_371, 1, x_369); -return x_371; +lean_ctor_set(x_373, 0, x_370); +lean_ctor_set(x_373, 1, x_371); +return x_373; } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_235); +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14723,32 +14730,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_372 = lean_ctor_get(x_238, 0); -lean_inc(x_372); -x_373 = lean_ctor_get(x_238, 1); -lean_inc(x_373); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - lean_ctor_release(x_238, 1); - x_374 = x_238; +x_374 = lean_ctor_get(x_240, 0); +lean_inc(x_374); +x_375 = lean_ctor_get(x_240, 1); +lean_inc(x_375); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_376 = x_240; } else { - lean_dec_ref(x_238); - x_374 = lean_box(0); + lean_dec_ref(x_240); + x_376 = lean_box(0); } -if (lean_is_scalar(x_374)) { - x_375 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_376)) { + x_377 = lean_alloc_ctor(1, 2, 0); } else { - x_375 = x_374; + x_377 = x_376; } -lean_ctor_set(x_375, 0, x_372); -lean_ctor_set(x_375, 1, x_373); -return x_375; +lean_ctor_set(x_377, 0, x_374); +lean_ctor_set(x_377, 1, x_375); +return x_377; } } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -lean_dec(x_235); +lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14764,33 +14771,33 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_376 = lean_ctor_get(x_236, 0); -lean_inc(x_376); -x_377 = lean_ctor_get(x_236, 1); -lean_inc(x_377); -if (lean_is_exclusive(x_236)) { - lean_ctor_release(x_236, 0); - lean_ctor_release(x_236, 1); - x_378 = x_236; +x_378 = lean_ctor_get(x_237, 0); +lean_inc(x_378); +x_379 = lean_ctor_get(x_237, 1); +lean_inc(x_379); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_380 = x_237; } else { - lean_dec_ref(x_236); - x_378 = lean_box(0); + lean_dec_ref(x_237); + x_380 = lean_box(0); } -if (lean_is_scalar(x_378)) { - x_379 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_380)) { + x_381 = lean_alloc_ctor(1, 2, 0); } else { - x_379 = x_378; + x_381 = x_380; } -lean_ctor_set(x_379, 0, x_376); -lean_ctor_set(x_379, 1, x_377); -return x_379; +lean_ctor_set(x_381, 0, x_378); +lean_ctor_set(x_381, 1, x_379); +return x_381; } } } } else { -uint8_t x_389; +uint8_t x_391; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -14807,23 +14814,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_389 = !lean_is_exclusive(x_21); -if (x_389 == 0) +x_391 = !lean_is_exclusive(x_21); +if (x_391 == 0) { return x_21; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_21, 0); -x_391 = lean_ctor_get(x_21, 1); -lean_inc(x_391); -lean_inc(x_390); +lean_object* x_392; lean_object* x_393; lean_object* x_394; +x_392 = lean_ctor_get(x_21, 0); +x_393 = lean_ctor_get(x_21, 1); +lean_inc(x_393); +lean_inc(x_392); lean_dec(x_21); -x_392 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_392, 0, x_390); -lean_ctor_set(x_392, 1, x_391); -return x_392; +x_394 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_394, 0, x_392); +lean_ctor_set(x_394, 1, x_393); +return x_394; } } } @@ -14856,148 +14863,149 @@ lean_dec(x_24); x_25 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); if (lean_obj_tag(x_25) == 1) { -lean_object* x_227; -x_227 = lean_ctor_get(x_25, 0); -lean_inc(x_227); -lean_ctor_set(x_21, 0, x_227); +lean_object* x_228; +x_228 = lean_ctor_get(x_25, 0); +lean_inc(x_228); +lean_ctor_set(x_21, 0, x_228); x_26 = x_21; -goto block_226; +goto block_227; } else { -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_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_free_object(x_21); lean_inc(x_25); -x_228 = l_Lean_MessageData_ofExpr(x_25); -x_229 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_230 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_230, 0, x_229); -lean_ctor_set(x_230, 1, x_228); -x_231 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; -x_232 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_232, 0, x_230); -lean_ctor_set(x_232, 1, x_231); -x_233 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_232, x_14, x_15, x_16, x_17, x_23); -x_26 = x_233; -goto block_226; -} -block_226: +x_229 = l_Lean_MessageData_ofExpr(x_25); +x_230 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_231 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_231, 0, x_230); +lean_ctor_set(x_231, 1, x_229); +x_232 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_233 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_233, 0, x_231); +lean_ctor_set(x_233, 1, x_232); +x_234 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_233, x_14, x_15, x_16, x_17, x_23); +x_26 = x_234; +goto block_227; +} +block_227: { if (lean_obj_tag(x_26) == 0) { -lean_object* x_27; lean_object* x_28; +lean_object* x_27; uint8_t x_28; lean_object* x_29; x_27 = lean_ctor_get(x_26, 1); lean_inc(x_27); lean_dec(x_26); +x_28 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_28 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); -if (lean_obj_tag(x_28) == 0) +x_29 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_28, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_27); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_29; lean_object* x_30; -x_29 = lean_ctor_get(x_28, 1); -lean_inc(x_29); -lean_dec(x_28); +lean_object* x_30; lean_object* x_31; +x_30 = lean_ctor_get(x_29, 1); +lean_inc(x_30); +lean_dec(x_29); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_25); -x_30 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, x_29); -if (lean_obj_tag(x_30) == 0) +x_31 = lean_infer_type(x_25, x_14, x_15, x_16, x_17, 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_object* x_32; lean_object* x_33; lean_object* x_34; +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); -lean_dec(x_30); +x_33 = lean_ctor_get(x_31, 1); +lean_inc(x_33); +lean_dec(x_31); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_33 = l_Lean_Meta_whnfD(x_31, x_14, x_15, x_16, x_17, x_32); -if (lean_obj_tag(x_33) == 0) +x_34 = l_Lean_Meta_whnfD(x_32, x_14, x_15, x_16, x_17, x_33); +if (lean_obj_tag(x_34) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_34 = lean_ctor_get(x_33, 0); -lean_inc(x_34); -x_35 = lean_ctor_get(x_33, 1); +lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; +x_35 = lean_ctor_get(x_34, 0); lean_inc(x_35); -lean_dec(x_33); -x_36 = l_Lean_Expr_getAppFn(x_34); -if (lean_obj_tag(x_36) == 4) +x_36 = lean_ctor_get(x_34, 1); +lean_inc(x_36); +lean_dec(x_34); +x_37 = l_Lean_Expr_getAppFn(x_35); +if (lean_obj_tag(x_37) == 4) { -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_ctor_get(x_36, 0); -lean_inc(x_110); -lean_dec(x_36); -x_111 = lean_st_ref_get(x_17, x_35); -x_112 = lean_ctor_get(x_111, 0); -lean_inc(x_112); -x_113 = lean_ctor_get(x_111, 1); +lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; +x_111 = lean_ctor_get(x_37, 0); +lean_inc(x_111); +lean_dec(x_37); +x_112 = lean_st_ref_get(x_17, x_36); +x_113 = lean_ctor_get(x_112, 0); lean_inc(x_113); -lean_dec(x_111); -x_114 = lean_ctor_get(x_112, 0); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); lean_dec(x_112); -x_115 = lean_environment_find(x_114, x_110); -if (lean_obj_tag(x_115) == 0) +x_115 = lean_ctor_get(x_113, 0); +lean_inc(x_115); +lean_dec(x_113); +x_116 = lean_environment_find(x_115, x_111); +if (lean_obj_tag(x_116) == 0) { -lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; lean_object* x_125; +lean_object* x_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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_116 = l_Lean_MessageData_ofExpr(x_25); -x_117 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_118 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_118, 0, x_117); -lean_ctor_set(x_118, 1, x_116); -x_119 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_120 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_120, 0, x_118); -lean_ctor_set(x_120, 1, x_119); -x_121 = l_Lean_MessageData_ofExpr(x_34); -x_122 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_122, 0, x_120); -lean_ctor_set(x_122, 1, x_121); -x_123 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_124 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_124, 0, x_122); -lean_ctor_set(x_124, 1, x_123); +x_117 = l_Lean_MessageData_ofExpr(x_25); +x_118 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_119 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_117); +x_120 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_121 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_121, 0, x_119); +lean_ctor_set(x_121, 1, x_120); +x_122 = l_Lean_MessageData_ofExpr(x_35); +x_123 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_123, 0, x_121); +lean_ctor_set(x_123, 1, x_122); +x_124 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_125 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_125, 0, x_123); +lean_ctor_set(x_125, 1, x_124); lean_inc(x_12); -x_125 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_124, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_125; -goto block_109; +x_126 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_125, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_126; +goto block_110; } else { -lean_object* x_126; -x_126 = lean_ctor_get(x_115, 0); -lean_inc(x_126); -lean_dec(x_115); -switch (lean_obj_tag(x_126)) { +lean_object* x_127; +x_127 = lean_ctor_get(x_116, 0); +lean_inc(x_127); +lean_dec(x_116); +switch (lean_obj_tag(x_127)) { case 4: { -lean_object* x_127; lean_object* x_128; uint8_t x_144; lean_object* x_145; +lean_object* x_128; lean_object* x_129; uint8_t x_145; lean_object* x_146; lean_dec(x_1); -x_127 = lean_ctor_get(x_126, 0); -lean_inc(x_127); -lean_dec(x_126); -x_144 = lean_ctor_get_uint8(x_127, sizeof(void*)*1); +x_128 = lean_ctor_get(x_127, 0); +lean_inc(x_128); lean_dec(x_127); -x_145 = lean_box(x_144); -if (lean_obj_tag(x_145) == 0) +x_145 = lean_ctor_get_uint8(x_128, sizeof(void*)*1); +lean_dec(x_128); +x_146 = lean_box(x_145); +if (lean_obj_tag(x_146) == 0) { -lean_object* x_146; lean_object* x_147; -lean_dec(x_34); -x_146 = lean_box(0); +lean_object* x_147; lean_object* x_148; +lean_dec(x_35); +x_147 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -15005,357 +15013,357 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_147 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_146, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_147; -goto block_109; +x_148 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_25, x_10, x_3, x_11, x_4, x_147, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_148; +goto block_110; } else { -lean_object* x_148; -lean_dec(x_145); +lean_object* x_149; +lean_dec(x_146); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_148 = lean_box(0); -x_128 = x_148; -goto block_143; +x_149 = lean_box(0); +x_129 = x_149; +goto block_144; } -block_143: +block_144: { -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; uint8_t x_139; -lean_dec(x_128); -x_129 = l_Lean_MessageData_ofExpr(x_25); -x_130 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_131 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_131, 0, x_130); -lean_ctor_set(x_131, 1, x_129); -x_132 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_133 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_133, 0, x_131); -lean_ctor_set(x_133, 1, x_132); -x_134 = l_Lean_MessageData_ofExpr(x_34); -x_135 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_135, 0, x_133); -lean_ctor_set(x_135, 1, x_134); -x_136 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_137 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_137, 0, x_135); -lean_ctor_set(x_137, 1, x_136); +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; uint8_t x_140; +lean_dec(x_129); +x_130 = l_Lean_MessageData_ofExpr(x_25); +x_131 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_132 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_132, 0, x_131); +lean_ctor_set(x_132, 1, x_130); +x_133 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_134 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_134, 0, x_132); +lean_ctor_set(x_134, 1, x_133); +x_135 = l_Lean_MessageData_ofExpr(x_35); +x_136 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_136, 0, x_134); +lean_ctor_set(x_136, 1, x_135); +x_137 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_138 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_138, 0, x_136); +lean_ctor_set(x_138, 1, x_137); lean_inc(x_12); -x_138 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_137, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_139 = !lean_is_exclusive(x_138); -if (x_139 == 0) +x_139 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_138, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_140 = !lean_is_exclusive(x_139); +if (x_140 == 0) { -x_37 = x_138; -goto block_109; +x_38 = x_139; +goto block_110; } else { -lean_object* x_140; lean_object* x_141; lean_object* x_142; -x_140 = lean_ctor_get(x_138, 0); -x_141 = lean_ctor_get(x_138, 1); +lean_object* x_141; lean_object* x_142; lean_object* x_143; +x_141 = lean_ctor_get(x_139, 0); +x_142 = lean_ctor_get(x_139, 1); +lean_inc(x_142); lean_inc(x_141); -lean_inc(x_140); -lean_dec(x_138); -x_142 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_142, 0, x_140); -lean_ctor_set(x_142, 1, x_141); -x_37 = x_142; -goto block_109; +lean_dec(x_139); +x_143 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_143, 0, x_141); +lean_ctor_set(x_143, 1, x_142); +x_38 = x_143; +goto block_110; } } } case 5: { -lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; -lean_dec(x_34); +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_35); lean_dec(x_11); lean_dec(x_3); -x_149 = lean_ctor_get(x_126, 0); -lean_inc(x_149); -lean_dec(x_126); -x_150 = lean_ctor_get(x_149, 1); +x_150 = lean_ctor_get(x_127, 0); lean_inc(x_150); -x_151 = lean_ctor_get(x_149, 4); +lean_dec(x_127); +x_151 = lean_ctor_get(x_150, 1); lean_inc(x_151); -lean_dec(x_149); -x_152 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_153 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +x_152 = lean_ctor_get(x_150, 4); +lean_inc(x_152); +lean_dec(x_150); +x_153 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_154 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_154 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_150, x_153, x_151, x_152, x_14, x_15, x_16, x_17, x_113); -lean_dec(x_150); -if (lean_obj_tag(x_154) == 0) +x_155 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_151, x_154, x_152, x_153, x_14, x_15, x_16, x_17, x_114); +lean_dec(x_151); +if (lean_obj_tag(x_155) == 0) { -lean_object* x_155; lean_object* x_156; uint8_t x_157; -x_155 = lean_ctor_get(x_154, 0); -lean_inc(x_155); -x_156 = lean_ctor_get(x_154, 1); +lean_object* x_156; lean_object* x_157; uint8_t x_158; +x_156 = lean_ctor_get(x_155, 0); lean_inc(x_156); -lean_dec(x_154); -x_157 = !lean_is_exclusive(x_155); -if (x_157 == 0) -{ -lean_object* x_158; lean_object* x_159; lean_object* x_160; uint8_t x_161; lean_object* x_162; -x_158 = lean_ctor_get(x_155, 0); -x_159 = lean_ctor_get(x_155, 1); -x_160 = l_Lean_Expr_fvarId_x21(x_25); -x_161 = 1; +x_157 = lean_ctor_get(x_155, 1); +lean_inc(x_157); +lean_dec(x_155); +x_158 = !lean_is_exclusive(x_156); +if (x_158 == 0) +{ +lean_object* x_159; lean_object* x_160; lean_object* x_161; uint8_t x_162; lean_object* x_163; +x_159 = lean_ctor_get(x_156, 0); +x_160 = lean_ctor_get(x_156, 1); +x_161 = l_Lean_Expr_fvarId_x21(x_25); +x_162 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_162 = l_Lean_Meta_Cases_cases(x_10, x_160, x_158, x_161, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_162) == 0) +x_163 = l_Lean_Meta_Cases_cases(x_10, x_161, x_159, x_162, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_163) == 0) { -uint8_t x_163; -x_163 = !lean_is_exclusive(x_162); -if (x_163 == 0) +uint8_t x_164; +x_164 = !lean_is_exclusive(x_163); +if (x_164 == 0) { -lean_object* x_164; -x_164 = lean_ctor_get(x_162, 0); -lean_ctor_set(x_155, 1, x_164); -lean_ctor_set(x_155, 0, x_159); -lean_ctor_set(x_162, 0, x_155); -x_37 = x_162; -goto block_109; +lean_object* x_165; +x_165 = lean_ctor_get(x_163, 0); +lean_ctor_set(x_156, 1, x_165); +lean_ctor_set(x_156, 0, x_160); +lean_ctor_set(x_163, 0, x_156); +x_38 = x_163; +goto block_110; } else { -lean_object* x_165; lean_object* x_166; lean_object* x_167; -x_165 = lean_ctor_get(x_162, 0); -x_166 = lean_ctor_get(x_162, 1); +lean_object* x_166; lean_object* x_167; lean_object* x_168; +x_166 = lean_ctor_get(x_163, 0); +x_167 = lean_ctor_get(x_163, 1); +lean_inc(x_167); lean_inc(x_166); -lean_inc(x_165); -lean_dec(x_162); -lean_ctor_set(x_155, 1, x_165); -lean_ctor_set(x_155, 0, x_159); -x_167 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_167, 0, x_155); -lean_ctor_set(x_167, 1, x_166); -x_37 = x_167; -goto block_109; +lean_dec(x_163); +lean_ctor_set(x_156, 1, x_166); +lean_ctor_set(x_156, 0, x_160); +x_168 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_168, 0, x_156); +lean_ctor_set(x_168, 1, x_167); +x_38 = x_168; +goto block_110; } } else { -uint8_t x_168; -lean_free_object(x_155); -lean_dec(x_159); -x_168 = !lean_is_exclusive(x_162); -if (x_168 == 0) +uint8_t x_169; +lean_free_object(x_156); +lean_dec(x_160); +x_169 = !lean_is_exclusive(x_163); +if (x_169 == 0) { -x_37 = x_162; -goto block_109; +x_38 = x_163; +goto block_110; } else { -lean_object* x_169; lean_object* x_170; lean_object* x_171; -x_169 = lean_ctor_get(x_162, 0); -x_170 = lean_ctor_get(x_162, 1); +lean_object* x_170; lean_object* x_171; lean_object* x_172; +x_170 = lean_ctor_get(x_163, 0); +x_171 = lean_ctor_get(x_163, 1); +lean_inc(x_171); lean_inc(x_170); -lean_inc(x_169); -lean_dec(x_162); -x_171 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_171, 0, x_169); -lean_ctor_set(x_171, 1, x_170); -x_37 = x_171; -goto block_109; +lean_dec(x_163); +x_172 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_172, 0, x_170); +lean_ctor_set(x_172, 1, x_171); +x_38 = x_172; +goto block_110; } } } else { -lean_object* x_172; lean_object* x_173; lean_object* x_174; uint8_t x_175; lean_object* x_176; -x_172 = lean_ctor_get(x_155, 0); -x_173 = lean_ctor_get(x_155, 1); +lean_object* x_173; lean_object* x_174; lean_object* x_175; uint8_t x_176; lean_object* x_177; +x_173 = lean_ctor_get(x_156, 0); +x_174 = lean_ctor_get(x_156, 1); +lean_inc(x_174); lean_inc(x_173); -lean_inc(x_172); -lean_dec(x_155); -x_174 = l_Lean_Expr_fvarId_x21(x_25); -x_175 = 1; +lean_dec(x_156); +x_175 = l_Lean_Expr_fvarId_x21(x_25); +x_176 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_176 = l_Lean_Meta_Cases_cases(x_10, x_174, x_172, x_175, x_14, x_15, x_16, x_17, x_156); -if (lean_obj_tag(x_176) == 0) +x_177 = l_Lean_Meta_Cases_cases(x_10, x_175, x_173, x_176, x_14, x_15, x_16, x_17, x_157); +if (lean_obj_tag(x_177) == 0) { -lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; -x_177 = lean_ctor_get(x_176, 0); -lean_inc(x_177); -x_178 = lean_ctor_get(x_176, 1); +lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; +x_178 = lean_ctor_get(x_177, 0); lean_inc(x_178); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_179 = x_176; +x_179 = lean_ctor_get(x_177, 1); +lean_inc(x_179); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_180 = x_177; } else { - lean_dec_ref(x_176); - x_179 = lean_box(0); + lean_dec_ref(x_177); + x_180 = lean_box(0); } -x_180 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_180, 0, x_173); -lean_ctor_set(x_180, 1, x_177); -if (lean_is_scalar(x_179)) { - x_181 = lean_alloc_ctor(0, 2, 0); +x_181 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_181, 0, x_174); +lean_ctor_set(x_181, 1, x_178); +if (lean_is_scalar(x_180)) { + x_182 = lean_alloc_ctor(0, 2, 0); } else { - x_181 = x_179; + x_182 = x_180; } -lean_ctor_set(x_181, 0, x_180); -lean_ctor_set(x_181, 1, x_178); -x_37 = x_181; -goto block_109; +lean_ctor_set(x_182, 0, x_181); +lean_ctor_set(x_182, 1, x_179); +x_38 = x_182; +goto block_110; } else { -lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; -lean_dec(x_173); -x_182 = lean_ctor_get(x_176, 0); -lean_inc(x_182); -x_183 = lean_ctor_get(x_176, 1); +lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; +lean_dec(x_174); +x_183 = lean_ctor_get(x_177, 0); lean_inc(x_183); -if (lean_is_exclusive(x_176)) { - lean_ctor_release(x_176, 0); - lean_ctor_release(x_176, 1); - x_184 = x_176; +x_184 = lean_ctor_get(x_177, 1); +lean_inc(x_184); +if (lean_is_exclusive(x_177)) { + lean_ctor_release(x_177, 0); + lean_ctor_release(x_177, 1); + x_185 = x_177; } else { - lean_dec_ref(x_176); - x_184 = lean_box(0); + lean_dec_ref(x_177); + x_185 = lean_box(0); } -if (lean_is_scalar(x_184)) { - x_185 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_185)) { + x_186 = lean_alloc_ctor(1, 2, 0); } else { - x_185 = x_184; + x_186 = x_185; } -lean_ctor_set(x_185, 0, x_182); -lean_ctor_set(x_185, 1, x_183); -x_37 = x_185; -goto block_109; +lean_ctor_set(x_186, 0, x_183); +lean_ctor_set(x_186, 1, x_184); +x_38 = x_186; +goto block_110; } } } else { -uint8_t x_186; +uint8_t x_187; lean_dec(x_25); lean_dec(x_10); -x_186 = !lean_is_exclusive(x_154); -if (x_186 == 0) +x_187 = !lean_is_exclusive(x_155); +if (x_187 == 0) { -x_37 = x_154; -goto block_109; +x_38 = x_155; +goto block_110; } else { -lean_object* x_187; lean_object* x_188; lean_object* x_189; -x_187 = lean_ctor_get(x_154, 0); -x_188 = lean_ctor_get(x_154, 1); +lean_object* x_188; lean_object* x_189; lean_object* x_190; +x_188 = lean_ctor_get(x_155, 0); +x_189 = lean_ctor_get(x_155, 1); +lean_inc(x_189); lean_inc(x_188); -lean_inc(x_187); -lean_dec(x_154); -x_189 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_189, 0, x_187); -lean_ctor_set(x_189, 1, x_188); -x_37 = x_189; -goto block_109; +lean_dec(x_155); +x_190 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_190, 0, x_188); +lean_ctor_set(x_190, 1, x_189); +x_38 = x_190; +goto block_110; } } } default: { -lean_object* x_190; lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; -lean_dec(x_126); +lean_object* x_191; lean_object* x_192; lean_object* x_193; lean_object* x_194; lean_object* x_195; lean_object* x_196; lean_object* x_197; lean_object* x_198; lean_object* x_199; lean_object* x_200; +lean_dec(x_127); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_190 = l_Lean_MessageData_ofExpr(x_25); -x_191 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_192 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_192, 0, x_191); -lean_ctor_set(x_192, 1, x_190); -x_193 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_194 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_194, 0, x_192); -lean_ctor_set(x_194, 1, x_193); -x_195 = l_Lean_MessageData_ofExpr(x_34); -x_196 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_196, 0, x_194); -lean_ctor_set(x_196, 1, x_195); -x_197 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_198 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_198, 0, x_196); -lean_ctor_set(x_198, 1, x_197); +x_191 = l_Lean_MessageData_ofExpr(x_25); +x_192 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_193 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_193, 0, x_192); +lean_ctor_set(x_193, 1, x_191); +x_194 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_195 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_195, 0, x_193); +lean_ctor_set(x_195, 1, x_194); +x_196 = l_Lean_MessageData_ofExpr(x_35); +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 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_199 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_199, 0, x_197); +lean_ctor_set(x_199, 1, x_198); lean_inc(x_12); -x_199 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_198, x_12, x_13, x_14, x_15, x_16, x_17, x_113); -x_37 = x_199; -goto block_109; +x_200 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_199, x_12, x_13, x_14, x_15, x_16, x_17, x_114); +x_38 = x_200; +goto block_110; } } } } else { -lean_object* x_200; lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; -lean_dec(x_36); +lean_object* x_201; lean_object* x_202; lean_object* x_203; lean_object* x_204; lean_object* x_205; lean_object* x_206; lean_object* x_207; lean_object* x_208; lean_object* x_209; lean_object* x_210; +lean_dec(x_37); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_200 = l_Lean_MessageData_ofExpr(x_25); -x_201 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_202 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_202, 0, x_201); -lean_ctor_set(x_202, 1, x_200); -x_203 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; -x_204 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_204, 0, x_202); -lean_ctor_set(x_204, 1, x_203); -x_205 = l_Lean_MessageData_ofExpr(x_34); -x_206 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_206, 0, x_204); -lean_ctor_set(x_206, 1, x_205); -x_207 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; -x_208 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_208, 0, x_206); -lean_ctor_set(x_208, 1, x_207); +x_201 = l_Lean_MessageData_ofExpr(x_25); +x_202 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; +x_203 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_203, 0, x_202); +lean_ctor_set(x_203, 1, x_201); +x_204 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_205 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_205, 0, x_203); +lean_ctor_set(x_205, 1, x_204); +x_206 = l_Lean_MessageData_ofExpr(x_35); +x_207 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_207, 0, x_205); +lean_ctor_set(x_207, 1, x_206); +x_208 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_209 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_209, 0, x_207); +lean_ctor_set(x_209, 1, x_208); lean_inc(x_12); -x_209 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_208, x_12, x_13, x_14, x_15, x_16, x_17, x_35); -x_37 = x_209; -goto block_109; +x_210 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_209, x_12, x_13, x_14, x_15, x_16, x_17, x_36); +x_38 = x_210; +goto block_110; } -block_109: +block_110: { -if (lean_obj_tag(x_37) == 0) +if (lean_obj_tag(x_38) == 0) { -uint8_t x_38; -x_38 = !lean_is_exclusive(x_37); -if (x_38 == 0) +uint8_t x_39; +x_39 = !lean_is_exclusive(x_38); +if (x_39 == 0) { -lean_object* x_39; uint8_t x_40; -x_39 = lean_ctor_get(x_37, 0); -x_40 = !lean_is_exclusive(x_39); -if (x_40 == 0) +lean_object* x_40; uint8_t x_41; +x_40 = lean_ctor_get(x_38, 0); +x_41 = !lean_is_exclusive(x_40); +if (x_41 == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; uint8_t x_45; -x_41 = lean_ctor_get(x_37, 1); -x_42 = lean_ctor_get(x_39, 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_38, 1); +x_43 = lean_ctor_get(x_40, 1); lean_inc(x_6); -lean_ctor_set(x_39, 1, x_6); -x_43 = lean_array_get_size(x_42); -x_44 = lean_unsigned_to_nat(0u); -x_45 = lean_nat_dec_lt(x_44, x_43); -if (x_45 == 0) +lean_ctor_set(x_40, 1, x_6); +x_44 = lean_array_get_size(x_43); +x_45 = lean_unsigned_to_nat(0u); +x_46 = lean_nat_dec_lt(x_45, x_44); +if (x_46 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15365,18 +15373,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_46; -x_46 = lean_nat_dec_le(x_43, x_43); -if (x_46 == 0) +uint8_t x_47; +x_47 = lean_nat_dec_le(x_44, x_44); +if (x_47 == 0) { +lean_dec(x_44); +lean_dec(x_40); lean_dec(x_43); -lean_dec(x_39); -lean_dec(x_42); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15386,70 +15394,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_47; size_t x_48; lean_object* x_49; -lean_free_object(x_37); +size_t x_48; size_t x_49; lean_object* x_50; +lean_free_object(x_38); lean_dec(x_6); -x_47 = 0; -x_48 = lean_usize_of_nat(x_43); +x_48 = 0; +x_49 = lean_usize_of_nat(x_44); +lean_dec(x_44); +x_50 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_43, x_48, x_49, x_40, x_12, x_13, x_14, x_15, x_16, x_17, x_42); lean_dec(x_43); -x_49 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_42, x_47, x_48, x_39, x_12, x_13, x_14, x_15, x_16, x_17, x_41); -lean_dec(x_42); -if (lean_obj_tag(x_49) == 0) +if (lean_obj_tag(x_50) == 0) { -uint8_t x_50; -x_50 = !lean_is_exclusive(x_49); -if (x_50 == 0) +uint8_t x_51; +x_51 = !lean_is_exclusive(x_50); +if (x_51 == 0) { -lean_object* x_51; lean_object* x_52; -x_51 = lean_ctor_get(x_49, 0); -x_52 = lean_ctor_get(x_51, 1); -lean_inc(x_52); -lean_dec(x_51); -lean_ctor_set(x_49, 0, x_52); -return x_49; +lean_object* x_52; lean_object* x_53; +x_52 = lean_ctor_get(x_50, 0); +x_53 = lean_ctor_get(x_52, 1); +lean_inc(x_53); +lean_dec(x_52); +lean_ctor_set(x_50, 0, x_53); +return x_50; } else { -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; -x_53 = lean_ctor_get(x_49, 0); -x_54 = lean_ctor_get(x_49, 1); -lean_inc(x_54); -lean_inc(x_53); -lean_dec(x_49); -x_55 = lean_ctor_get(x_53, 1); +lean_object* x_54; lean_object* x_55; lean_object* x_56; lean_object* x_57; +x_54 = lean_ctor_get(x_50, 0); +x_55 = lean_ctor_get(x_50, 1); lean_inc(x_55); -lean_dec(x_53); -x_56 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_56, 0, x_55); -lean_ctor_set(x_56, 1, x_54); -return x_56; +lean_inc(x_54); +lean_dec(x_50); +x_56 = lean_ctor_get(x_54, 1); +lean_inc(x_56); +lean_dec(x_54); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_56); +lean_ctor_set(x_57, 1, x_55); +return x_57; } } else { -uint8_t x_57; -x_57 = !lean_is_exclusive(x_49); -if (x_57 == 0) +uint8_t x_58; +x_58 = !lean_is_exclusive(x_50); +if (x_58 == 0) { -return x_49; +return x_50; } else { -lean_object* x_58; lean_object* x_59; lean_object* x_60; -x_58 = lean_ctor_get(x_49, 0); -x_59 = lean_ctor_get(x_49, 1); +lean_object* x_59; lean_object* x_60; lean_object* x_61; +x_59 = lean_ctor_get(x_50, 0); +x_60 = lean_ctor_get(x_50, 1); +lean_inc(x_60); lean_inc(x_59); -lean_inc(x_58); -lean_dec(x_49); -x_60 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_60, 0, x_58); -lean_ctor_set(x_60, 1, x_59); -return x_60; +lean_dec(x_50); +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; } } } @@ -15457,25 +15465,25 @@ return x_60; } else { -lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; uint8_t x_67; -x_61 = lean_ctor_get(x_37, 1); -x_62 = lean_ctor_get(x_39, 0); -x_63 = lean_ctor_get(x_39, 1); +lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; uint8_t x_68; +x_62 = lean_ctor_get(x_38, 1); +x_63 = lean_ctor_get(x_40, 0); +x_64 = lean_ctor_get(x_40, 1); +lean_inc(x_64); lean_inc(x_63); -lean_inc(x_62); -lean_dec(x_39); +lean_dec(x_40); lean_inc(x_6); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_62); -lean_ctor_set(x_64, 1, x_6); -x_65 = lean_array_get_size(x_63); -x_66 = lean_unsigned_to_nat(0u); -x_67 = lean_nat_dec_lt(x_66, x_65); -if (x_67 == 0) +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_6); +x_66 = lean_array_get_size(x_64); +x_67 = lean_unsigned_to_nat(0u); +x_68 = lean_nat_dec_lt(x_67, x_66); +if (x_68 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15485,18 +15493,18 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -uint8_t x_68; -x_68 = lean_nat_dec_le(x_65, x_65); -if (x_68 == 0) +uint8_t x_69; +x_69 = lean_nat_dec_le(x_66, x_66); +if (x_69 == 0) { +lean_dec(x_66); lean_dec(x_65); lean_dec(x_64); -lean_dec(x_63); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15506,69 +15514,69 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -lean_ctor_set(x_37, 0, x_6); -return x_37; +lean_ctor_set(x_38, 0, x_6); +return x_38; } else { -size_t x_69; size_t x_70; lean_object* x_71; -lean_free_object(x_37); +size_t x_70; size_t x_71; lean_object* x_72; +lean_free_object(x_38); lean_dec(x_6); -x_69 = 0; -x_70 = lean_usize_of_nat(x_65); -lean_dec(x_65); -x_71 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_63, x_69, x_70, x_64, x_12, x_13, x_14, x_15, x_16, x_17, x_61); -lean_dec(x_63); -if (lean_obj_tag(x_71) == 0) +x_70 = 0; +x_71 = lean_usize_of_nat(x_66); +lean_dec(x_66); +x_72 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_64, x_70, x_71, x_65, x_12, x_13, x_14, x_15, x_16, x_17, x_62); +lean_dec(x_64); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_72; lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; -x_72 = lean_ctor_get(x_71, 0); -lean_inc(x_72); -x_73 = lean_ctor_get(x_71, 1); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); 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_74 = lean_ctor_get(x_72, 1); +lean_inc(x_74); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; } else { - lean_dec_ref(x_71); - x_74 = lean_box(0); + lean_dec_ref(x_72); + x_75 = lean_box(0); } -x_75 = lean_ctor_get(x_72, 1); -lean_inc(x_75); -lean_dec(x_72); -if (lean_is_scalar(x_74)) { - x_76 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_ctor_get(x_73, 1); +lean_inc(x_76); +lean_dec(x_73); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_76 = x_74; + x_77 = x_75; } -lean_ctor_set(x_76, 0, x_75); -lean_ctor_set(x_76, 1, x_73); -return x_76; +lean_ctor_set(x_77, 0, x_76); +lean_ctor_set(x_77, 1, x_74); +return x_77; } else { -lean_object* x_77; lean_object* x_78; lean_object* x_79; lean_object* x_80; -x_77 = lean_ctor_get(x_71, 0); -lean_inc(x_77); -x_78 = lean_ctor_get(x_71, 1); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +x_78 = lean_ctor_get(x_72, 0); lean_inc(x_78); -if (lean_is_exclusive(x_71)) { - lean_ctor_release(x_71, 0); - lean_ctor_release(x_71, 1); - x_79 = x_71; +x_79 = lean_ctor_get(x_72, 1); +lean_inc(x_79); +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_80 = x_72; } else { - lean_dec_ref(x_71); - x_79 = lean_box(0); + lean_dec_ref(x_72); + x_80 = lean_box(0); } -if (lean_is_scalar(x_79)) { - x_80 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_80 = x_79; + x_81 = x_80; } -lean_ctor_set(x_80, 0, x_77); -lean_ctor_set(x_80, 1, x_78); -return x_80; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } @@ -15576,41 +15584,41 @@ 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; uint8_t x_89; -x_81 = lean_ctor_get(x_37, 0); -x_82 = lean_ctor_get(x_37, 1); -lean_inc(x_82); -lean_inc(x_81); -lean_dec(x_37); -x_83 = lean_ctor_get(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; lean_object* x_88; lean_object* x_89; uint8_t x_90; +x_82 = lean_ctor_get(x_38, 0); +x_83 = lean_ctor_get(x_38, 1); lean_inc(x_83); -x_84 = lean_ctor_get(x_81, 1); +lean_inc(x_82); +lean_dec(x_38); +x_84 = lean_ctor_get(x_82, 0); lean_inc(x_84); -if (lean_is_exclusive(x_81)) { - lean_ctor_release(x_81, 0); - lean_ctor_release(x_81, 1); - x_85 = x_81; +x_85 = lean_ctor_get(x_82, 1); +lean_inc(x_85); +if (lean_is_exclusive(x_82)) { + lean_ctor_release(x_82, 0); + lean_ctor_release(x_82, 1); + x_86 = x_82; } else { - lean_dec_ref(x_81); - x_85 = lean_box(0); + lean_dec_ref(x_82); + x_86 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_85)) { - x_86 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_86)) { + x_87 = lean_alloc_ctor(0, 2, 0); } else { - x_86 = x_85; + x_87 = x_86; } -lean_ctor_set(x_86, 0, x_83); -lean_ctor_set(x_86, 1, x_6); -x_87 = lean_array_get_size(x_84); -x_88 = lean_unsigned_to_nat(0u); -x_89 = lean_nat_dec_lt(x_88, x_87); -if (x_89 == 0) +lean_ctor_set(x_87, 0, x_84); +lean_ctor_set(x_87, 1, x_6); +x_88 = lean_array_get_size(x_85); +x_89 = lean_unsigned_to_nat(0u); +x_90 = lean_nat_dec_lt(x_89, x_88); +if (x_90 == 0) { -lean_object* x_90; +lean_object* x_91; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15620,21 +15628,21 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_90 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_90, 0, x_6); -lean_ctor_set(x_90, 1, x_82); -return x_90; +x_91 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_91, 0, x_6); +lean_ctor_set(x_91, 1, x_83); +return x_91; } else { -uint8_t x_91; -x_91 = lean_nat_dec_le(x_87, x_87); -if (x_91 == 0) +uint8_t x_92; +x_92 = lean_nat_dec_le(x_88, x_88); +if (x_92 == 0) { -lean_object* x_92; +lean_object* x_93; +lean_dec(x_88); lean_dec(x_87); -lean_dec(x_86); -lean_dec(x_84); +lean_dec(x_85); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15644,70 +15652,70 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -x_92 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_92, 0, x_6); -lean_ctor_set(x_92, 1, x_82); -return x_92; +x_93 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_93, 0, x_6); +lean_ctor_set(x_93, 1, x_83); +return x_93; } else { -size_t x_93; size_t x_94; lean_object* x_95; +size_t x_94; size_t x_95; lean_object* x_96; lean_dec(x_6); -x_93 = 0; -x_94 = lean_usize_of_nat(x_87); -lean_dec(x_87); -x_95 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_84, x_93, x_94, x_86, x_12, x_13, x_14, x_15, x_16, x_17, x_82); -lean_dec(x_84); -if (lean_obj_tag(x_95) == 0) +x_94 = 0; +x_95 = lean_usize_of_nat(x_88); +lean_dec(x_88); +x_96 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_85, x_94, x_95, x_87, x_12, x_13, x_14, x_15, x_16, x_17, x_83); +lean_dec(x_85); +if (lean_obj_tag(x_96) == 0) { -lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; -x_96 = lean_ctor_get(x_95, 0); -lean_inc(x_96); -x_97 = lean_ctor_get(x_95, 1); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; +x_97 = lean_ctor_get(x_96, 0); lean_inc(x_97); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_98 = x_95; +x_98 = lean_ctor_get(x_96, 1); +lean_inc(x_98); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_99 = x_96; } else { - lean_dec_ref(x_95); - x_98 = lean_box(0); + lean_dec_ref(x_96); + x_99 = lean_box(0); } -x_99 = lean_ctor_get(x_96, 1); -lean_inc(x_99); -lean_dec(x_96); -if (lean_is_scalar(x_98)) { - x_100 = lean_alloc_ctor(0, 2, 0); +x_100 = lean_ctor_get(x_97, 1); +lean_inc(x_100); +lean_dec(x_97); +if (lean_is_scalar(x_99)) { + x_101 = lean_alloc_ctor(0, 2, 0); } else { - x_100 = x_98; + x_101 = x_99; } -lean_ctor_set(x_100, 0, x_99); -lean_ctor_set(x_100, 1, x_97); -return x_100; +lean_ctor_set(x_101, 0, x_100); +lean_ctor_set(x_101, 1, x_98); +return x_101; } else { -lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; -x_101 = lean_ctor_get(x_95, 0); -lean_inc(x_101); -x_102 = lean_ctor_get(x_95, 1); +lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; +x_102 = lean_ctor_get(x_96, 0); lean_inc(x_102); -if (lean_is_exclusive(x_95)) { - lean_ctor_release(x_95, 0); - lean_ctor_release(x_95, 1); - x_103 = x_95; +x_103 = lean_ctor_get(x_96, 1); +lean_inc(x_103); +if (lean_is_exclusive(x_96)) { + lean_ctor_release(x_96, 0); + lean_ctor_release(x_96, 1); + x_104 = x_96; } else { - lean_dec_ref(x_95); - x_103 = lean_box(0); + lean_dec_ref(x_96); + x_104 = lean_box(0); } -if (lean_is_scalar(x_103)) { - x_104 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_104)) { + x_105 = lean_alloc_ctor(1, 2, 0); } else { - x_104 = x_103; + x_105 = x_104; } -lean_ctor_set(x_104, 0, x_101); -lean_ctor_set(x_104, 1, x_102); -return x_104; +lean_ctor_set(x_105, 0, x_102); +lean_ctor_set(x_105, 1, x_103); +return x_105; } } } @@ -15715,7 +15723,7 @@ return x_104; } else { -uint8_t x_105; +uint8_t x_106; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -15726,30 +15734,30 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_105 = !lean_is_exclusive(x_37); -if (x_105 == 0) +x_106 = !lean_is_exclusive(x_38); +if (x_106 == 0) { -return x_37; +return x_38; } else { -lean_object* x_106; lean_object* x_107; lean_object* x_108; -x_106 = lean_ctor_get(x_37, 0); -x_107 = lean_ctor_get(x_37, 1); +lean_object* x_107; lean_object* x_108; lean_object* x_109; +x_107 = lean_ctor_get(x_38, 0); +x_108 = lean_ctor_get(x_38, 1); +lean_inc(x_108); lean_inc(x_107); -lean_inc(x_106); -lean_dec(x_37); -x_108 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_108, 0, x_106); -lean_ctor_set(x_108, 1, x_107); -return x_108; +lean_dec(x_38); +x_109 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_109, 0, x_107); +lean_ctor_set(x_109, 1, x_108); +return x_109; } } } } else { -uint8_t x_210; +uint8_t x_211; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -15766,29 +15774,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_210 = !lean_is_exclusive(x_33); -if (x_210 == 0) +x_211 = !lean_is_exclusive(x_34); +if (x_211 == 0) { -return x_33; +return x_34; } else { -lean_object* x_211; lean_object* x_212; lean_object* x_213; -x_211 = lean_ctor_get(x_33, 0); -x_212 = lean_ctor_get(x_33, 1); +lean_object* x_212; lean_object* x_213; lean_object* x_214; +x_212 = lean_ctor_get(x_34, 0); +x_213 = lean_ctor_get(x_34, 1); +lean_inc(x_213); lean_inc(x_212); -lean_inc(x_211); -lean_dec(x_33); -x_213 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_213, 0, x_211); -lean_ctor_set(x_213, 1, x_212); -return x_213; +lean_dec(x_34); +x_214 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_214, 0, x_212); +lean_ctor_set(x_214, 1, x_213); +return x_214; } } } else { -uint8_t x_214; +uint8_t x_215; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -15805,29 +15813,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_214 = !lean_is_exclusive(x_30); -if (x_214 == 0) +x_215 = !lean_is_exclusive(x_31); +if (x_215 == 0) { -return x_30; +return x_31; } else { -lean_object* x_215; lean_object* x_216; lean_object* x_217; -x_215 = lean_ctor_get(x_30, 0); -x_216 = lean_ctor_get(x_30, 1); +lean_object* x_216; lean_object* x_217; lean_object* x_218; +x_216 = lean_ctor_get(x_31, 0); +x_217 = lean_ctor_get(x_31, 1); +lean_inc(x_217); lean_inc(x_216); -lean_inc(x_215); -lean_dec(x_30); -x_217 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_217, 0, x_215); -lean_ctor_set(x_217, 1, x_216); -return x_217; +lean_dec(x_31); +x_218 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_218, 0, x_216); +lean_ctor_set(x_218, 1, x_217); +return x_218; } } } else { -uint8_t x_218; +uint8_t x_219; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -15844,29 +15852,29 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_218 = !lean_is_exclusive(x_28); -if (x_218 == 0) +x_219 = !lean_is_exclusive(x_29); +if (x_219 == 0) { -return x_28; +return x_29; } else { -lean_object* x_219; lean_object* x_220; lean_object* x_221; -x_219 = lean_ctor_get(x_28, 0); -x_220 = lean_ctor_get(x_28, 1); +lean_object* x_220; lean_object* x_221; lean_object* x_222; +x_220 = lean_ctor_get(x_29, 0); +x_221 = lean_ctor_get(x_29, 1); +lean_inc(x_221); lean_inc(x_220); -lean_inc(x_219); -lean_dec(x_28); -x_221 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_221, 0, x_219); -lean_ctor_set(x_221, 1, x_220); -return x_221; +lean_dec(x_29); +x_222 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_222, 0, x_220); +lean_ctor_set(x_222, 1, x_221); +return x_222; } } } else { -uint8_t x_222; +uint8_t x_223; lean_dec(x_25); lean_dec(x_17); lean_dec(x_16); @@ -15883,179 +15891,180 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_222 = !lean_is_exclusive(x_26); -if (x_222 == 0) +x_223 = !lean_is_exclusive(x_26); +if (x_223 == 0) { return x_26; } else { -lean_object* x_223; lean_object* x_224; lean_object* x_225; -x_223 = lean_ctor_get(x_26, 0); -x_224 = lean_ctor_get(x_26, 1); +lean_object* x_224; lean_object* x_225; lean_object* x_226; +x_224 = lean_ctor_get(x_26, 0); +x_225 = lean_ctor_get(x_26, 1); +lean_inc(x_225); lean_inc(x_224); -lean_inc(x_223); lean_dec(x_26); -x_225 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_225, 0, x_223); -lean_ctor_set(x_225, 1, x_224); -return x_225; +x_226 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_226, 0, x_224); +lean_ctor_set(x_226, 1, x_225); +return x_226; } } } } else { -lean_object* x_234; lean_object* x_235; lean_object* x_236; -x_234 = lean_ctor_get(x_21, 1); -lean_inc(x_234); +lean_object* x_235; lean_object* x_236; lean_object* x_237; +x_235 = lean_ctor_get(x_21, 1); +lean_inc(x_235); lean_dec(x_21); -x_235 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); -if (lean_obj_tag(x_235) == 1) +x_236 = l_Lean_Meta_FVarSubst_apply(x_4, x_5); +if (lean_obj_tag(x_236) == 1) { -lean_object* x_381; lean_object* x_382; -x_381 = lean_ctor_get(x_235, 0); -lean_inc(x_381); -x_382 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_382, 0, x_381); -lean_ctor_set(x_382, 1, x_234); -x_236 = x_382; -goto block_380; +lean_object* x_383; lean_object* x_384; +x_383 = lean_ctor_get(x_236, 0); +lean_inc(x_383); +x_384 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_384, 0, x_383); +lean_ctor_set(x_384, 1, x_235); +x_237 = x_384; +goto block_382; } else { -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_inc(x_235); -x_383 = l_Lean_MessageData_ofExpr(x_235); -x_384 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_385 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_385, 0, x_384); -lean_ctor_set(x_385, 1, x_383); -x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +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_inc(x_236); +x_385 = l_Lean_MessageData_ofExpr(x_236); +x_386 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_387 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_387, 0, x_385); -lean_ctor_set(x_387, 1, x_386); -x_388 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_387, x_14, x_15, x_16, x_17, x_234); -x_236 = x_388; -goto block_380; -} -block_380: -{ -if (lean_obj_tag(x_236) == 0) -{ -lean_object* x_237; lean_object* x_238; -x_237 = lean_ctor_get(x_236, 1); -lean_inc(x_237); -lean_dec(x_236); +lean_ctor_set(x_387, 0, x_386); +lean_ctor_set(x_387, 1, x_385); +x_388 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__4; +x_389 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_389, 0, x_387); +lean_ctor_set(x_389, 1, x_388); +x_390 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__1(x_389, x_14, x_15, x_16, x_17, x_235); +x_237 = x_390; +goto block_382; +} +block_382: +{ +if (lean_obj_tag(x_237) == 0) +{ +lean_object* x_238; uint8_t x_239; lean_object* x_240; +x_238 = lean_ctor_get(x_237, 1); +lean_inc(x_238); +lean_dec(x_237); +x_239 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); -x_238 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_20, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_237); -if (lean_obj_tag(x_238) == 0) -{ -lean_object* x_239; lean_object* x_240; -x_239 = lean_ctor_get(x_238, 1); -lean_inc(x_239); -lean_dec(x_238); -lean_inc(x_17); -lean_inc(x_16); -lean_inc(x_15); -lean_inc(x_14); -lean_inc(x_235); -x_240 = lean_infer_type(x_235, x_14, x_15, x_16, x_17, x_239); +x_240 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_239, x_20, x_12, x_13, x_14, x_15, x_16, x_17, x_238); if (lean_obj_tag(x_240) == 0) { -lean_object* x_241; lean_object* x_242; lean_object* x_243; -x_241 = lean_ctor_get(x_240, 0); +lean_object* x_241; lean_object* x_242; +x_241 = lean_ctor_get(x_240, 1); lean_inc(x_241); -x_242 = lean_ctor_get(x_240, 1); -lean_inc(x_242); lean_dec(x_240); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_243 = l_Lean_Meta_whnfD(x_241, x_14, x_15, x_16, x_17, x_242); -if (lean_obj_tag(x_243) == 0) +lean_inc(x_236); +x_242 = lean_infer_type(x_236, x_14, x_15, x_16, x_17, x_241); +if (lean_obj_tag(x_242) == 0) { -lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; -x_244 = lean_ctor_get(x_243, 0); +lean_object* x_243; lean_object* x_244; lean_object* x_245; +x_243 = lean_ctor_get(x_242, 0); +lean_inc(x_243); +x_244 = lean_ctor_get(x_242, 1); lean_inc(x_244); -x_245 = lean_ctor_get(x_243, 1); -lean_inc(x_245); -lean_dec(x_243); -x_246 = l_Lean_Expr_getAppFn(x_244); -if (lean_obj_tag(x_246) == 4) -{ -lean_object* x_278; lean_object* x_279; lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; -x_278 = lean_ctor_get(x_246, 0); -lean_inc(x_278); -lean_dec(x_246); -x_279 = lean_st_ref_get(x_17, x_245); -x_280 = lean_ctor_get(x_279, 0); +lean_dec(x_242); +lean_inc(x_17); +lean_inc(x_16); +lean_inc(x_15); +lean_inc(x_14); +x_245 = l_Lean_Meta_whnfD(x_243, x_14, x_15, x_16, x_17, x_244); +if (lean_obj_tag(x_245) == 0) +{ +lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; +x_246 = lean_ctor_get(x_245, 0); +lean_inc(x_246); +x_247 = lean_ctor_get(x_245, 1); +lean_inc(x_247); +lean_dec(x_245); +x_248 = l_Lean_Expr_getAppFn(x_246); +if (lean_obj_tag(x_248) == 4) +{ +lean_object* x_280; lean_object* x_281; lean_object* x_282; lean_object* x_283; lean_object* x_284; lean_object* x_285; +x_280 = lean_ctor_get(x_248, 0); lean_inc(x_280); -x_281 = lean_ctor_get(x_279, 1); -lean_inc(x_281); -lean_dec(x_279); -x_282 = lean_ctor_get(x_280, 0); +lean_dec(x_248); +x_281 = lean_st_ref_get(x_17, x_247); +x_282 = lean_ctor_get(x_281, 0); lean_inc(x_282); -lean_dec(x_280); -x_283 = lean_environment_find(x_282, x_278); -if (lean_obj_tag(x_283) == 0) -{ -lean_object* x_284; lean_object* x_285; lean_object* x_286; lean_object* x_287; lean_object* x_288; lean_object* x_289; lean_object* x_290; lean_object* x_291; lean_object* x_292; lean_object* x_293; +x_283 = lean_ctor_get(x_281, 1); +lean_inc(x_283); +lean_dec(x_281); +x_284 = lean_ctor_get(x_282, 0); +lean_inc(x_284); +lean_dec(x_282); +x_285 = lean_environment_find(x_284, x_280); +if (lean_obj_tag(x_285) == 0) +{ +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_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_284 = l_Lean_MessageData_ofExpr(x_235); -x_285 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_286 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_286, 0, x_285); -lean_ctor_set(x_286, 1, x_284); -x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_286 = l_Lean_MessageData_ofExpr(x_236); +x_287 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_288 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_288, 0, x_286); -lean_ctor_set(x_288, 1, x_287); -x_289 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_288, 0, x_287); +lean_ctor_set(x_288, 1, x_286); +x_289 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_290 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_290, 0, x_288); lean_ctor_set(x_290, 1, x_289); -x_291 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_291 = l_Lean_MessageData_ofExpr(x_246); x_292 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_292, 0, x_290); lean_ctor_set(x_292, 1, x_291); +x_293 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_294 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_294, 0, x_292); +lean_ctor_set(x_294, 1, x_293); lean_inc(x_12); -x_293 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_292, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_293; -goto block_277; +x_295 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_294, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_295; +goto block_279; } else { -lean_object* x_294; -x_294 = lean_ctor_get(x_283, 0); -lean_inc(x_294); -lean_dec(x_283); -switch (lean_obj_tag(x_294)) { +lean_object* x_296; +x_296 = lean_ctor_get(x_285, 0); +lean_inc(x_296); +lean_dec(x_285); +switch (lean_obj_tag(x_296)) { case 4: { -lean_object* x_295; lean_object* x_296; uint8_t x_312; lean_object* x_313; +lean_object* x_297; lean_object* x_298; uint8_t x_314; lean_object* x_315; lean_dec(x_1); -x_295 = lean_ctor_get(x_294, 0); -lean_inc(x_295); -lean_dec(x_294); -x_312 = lean_ctor_get_uint8(x_295, sizeof(void*)*1); -lean_dec(x_295); -x_313 = lean_box(x_312); -if (lean_obj_tag(x_313) == 0) -{ -lean_object* x_314; lean_object* x_315; -lean_dec(x_244); -x_314 = lean_box(0); +x_297 = lean_ctor_get(x_296, 0); +lean_inc(x_297); +lean_dec(x_296); +x_314 = lean_ctor_get_uint8(x_297, sizeof(void*)*1); +lean_dec(x_297); +x_315 = lean_box(x_314); +if (lean_obj_tag(x_315) == 0) +{ +lean_object* x_316; lean_object* x_317; +lean_dec(x_246); +x_316 = lean_box(0); lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); @@ -16063,316 +16072,316 @@ lean_inc(x_14); lean_inc(x_13); lean_inc(x_12); lean_inc(x_4); -x_315 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_235, x_10, x_3, x_11, x_4, x_314, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_315; -goto block_277; +x_317 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__7(x_9, x_236, x_10, x_3, x_11, x_4, x_316, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_317; +goto block_279; } else { -lean_object* x_316; -lean_dec(x_313); +lean_object* x_318; +lean_dec(x_315); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); -x_316 = lean_box(0); -x_296 = x_316; -goto block_311; +x_318 = lean_box(0); +x_298 = x_318; +goto block_313; } -block_311: +block_313: { -lean_object* x_297; lean_object* x_298; lean_object* x_299; lean_object* x_300; lean_object* x_301; lean_object* x_302; lean_object* x_303; lean_object* x_304; lean_object* x_305; lean_object* x_306; lean_object* x_307; lean_object* x_308; lean_object* x_309; lean_object* x_310; -lean_dec(x_296); -x_297 = l_Lean_MessageData_ofExpr(x_235); -x_298 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_299 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_299, 0, x_298); -lean_ctor_set(x_299, 1, x_297); -x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +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_dec(x_298); +x_299 = l_Lean_MessageData_ofExpr(x_236); +x_300 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_301 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_301, 0, x_299); -lean_ctor_set(x_301, 1, x_300); -x_302 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_301, 0, x_300); +lean_ctor_set(x_301, 1, x_299); +x_302 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_303 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_303, 0, x_301); lean_ctor_set(x_303, 1, x_302); -x_304 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_304 = l_Lean_MessageData_ofExpr(x_246); x_305 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_305, 0, x_303); lean_ctor_set(x_305, 1, x_304); +x_306 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_307 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_307, 0, x_305); +lean_ctor_set(x_307, 1, x_306); lean_inc(x_12); -x_306 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_305, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_307 = lean_ctor_get(x_306, 0); -lean_inc(x_307); -x_308 = lean_ctor_get(x_306, 1); -lean_inc(x_308); -if (lean_is_exclusive(x_306)) { - lean_ctor_release(x_306, 0); - lean_ctor_release(x_306, 1); - x_309 = x_306; +x_308 = l_Lean_throwError___at_Lean_Elab_Term_synthesizeInstMVarCore___spec__3(x_307, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_309 = lean_ctor_get(x_308, 0); +lean_inc(x_309); +x_310 = lean_ctor_get(x_308, 1); +lean_inc(x_310); +if (lean_is_exclusive(x_308)) { + lean_ctor_release(x_308, 0); + lean_ctor_release(x_308, 1); + x_311 = x_308; } else { - lean_dec_ref(x_306); - x_309 = lean_box(0); + lean_dec_ref(x_308); + x_311 = lean_box(0); } -if (lean_is_scalar(x_309)) { - x_310 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_311)) { + x_312 = lean_alloc_ctor(1, 2, 0); } else { - x_310 = x_309; + x_312 = x_311; } -lean_ctor_set(x_310, 0, x_307); -lean_ctor_set(x_310, 1, x_308); -x_247 = x_310; -goto block_277; +lean_ctor_set(x_312, 0, x_309); +lean_ctor_set(x_312, 1, x_310); +x_249 = x_312; +goto block_279; } } case 5: { -lean_object* x_317; lean_object* x_318; lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; -lean_dec(x_244); +lean_object* x_319; lean_object* x_320; lean_object* x_321; lean_object* x_322; lean_object* x_323; lean_object* x_324; +lean_dec(x_246); lean_dec(x_11); lean_dec(x_3); -x_317 = lean_ctor_get(x_294, 0); -lean_inc(x_317); -lean_dec(x_294); -x_318 = lean_ctor_get(x_317, 1); -lean_inc(x_318); -x_319 = lean_ctor_get(x_317, 4); +x_319 = lean_ctor_get(x_296, 0); lean_inc(x_319); -lean_dec(x_317); -x_320 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); -x_321 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; +lean_dec(x_296); +x_320 = lean_ctor_get(x_319, 1); +lean_inc(x_320); +x_321 = lean_ctor_get(x_319, 4); +lean_inc(x_321); +lean_dec(x_319); +x_322 = l_Lean_Elab_Tactic_RCases_RCasesPatt_asAlts(x_9); +x_323 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_322 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_318, x_321, x_319, x_320, x_14, x_15, x_16, x_17, x_281); -lean_dec(x_318); -if (lean_obj_tag(x_322) == 0) -{ -lean_object* x_323; lean_object* x_324; lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; uint8_t x_329; lean_object* x_330; -x_323 = lean_ctor_get(x_322, 0); -lean_inc(x_323); -x_324 = lean_ctor_get(x_322, 1); -lean_inc(x_324); -lean_dec(x_322); -x_325 = lean_ctor_get(x_323, 0); +x_324 = l_Lean_Elab_Tactic_RCases_processConstructors(x_1, x_320, x_323, x_321, x_322, x_14, x_15, x_16, x_17, x_283); +lean_dec(x_320); +if (lean_obj_tag(x_324) == 0) +{ +lean_object* x_325; lean_object* x_326; lean_object* x_327; lean_object* x_328; lean_object* x_329; lean_object* x_330; uint8_t x_331; lean_object* x_332; +x_325 = lean_ctor_get(x_324, 0); lean_inc(x_325); -x_326 = lean_ctor_get(x_323, 1); +x_326 = lean_ctor_get(x_324, 1); lean_inc(x_326); -if (lean_is_exclusive(x_323)) { - lean_ctor_release(x_323, 0); - lean_ctor_release(x_323, 1); - x_327 = x_323; -} else { - lean_dec_ref(x_323); - x_327 = lean_box(0); -} -x_328 = l_Lean_Expr_fvarId_x21(x_235); -x_329 = 1; +lean_dec(x_324); +x_327 = lean_ctor_get(x_325, 0); +lean_inc(x_327); +x_328 = lean_ctor_get(x_325, 1); +lean_inc(x_328); +if (lean_is_exclusive(x_325)) { + lean_ctor_release(x_325, 0); + lean_ctor_release(x_325, 1); + x_329 = x_325; +} else { + lean_dec_ref(x_325); + x_329 = lean_box(0); +} +x_330 = l_Lean_Expr_fvarId_x21(x_236); +x_331 = 1; lean_inc(x_17); lean_inc(x_16); lean_inc(x_15); lean_inc(x_14); -x_330 = l_Lean_Meta_Cases_cases(x_10, x_328, x_325, x_329, x_14, x_15, x_16, x_17, x_324); -if (lean_obj_tag(x_330) == 0) +x_332 = l_Lean_Meta_Cases_cases(x_10, x_330, x_327, x_331, x_14, x_15, x_16, x_17, x_326); +if (lean_obj_tag(x_332) == 0) { -lean_object* x_331; lean_object* x_332; lean_object* x_333; lean_object* x_334; lean_object* x_335; -x_331 = lean_ctor_get(x_330, 0); -lean_inc(x_331); -x_332 = lean_ctor_get(x_330, 1); -lean_inc(x_332); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_333 = x_330; +lean_object* x_333; lean_object* x_334; lean_object* x_335; lean_object* x_336; lean_object* x_337; +x_333 = lean_ctor_get(x_332, 0); +lean_inc(x_333); +x_334 = lean_ctor_get(x_332, 1); +lean_inc(x_334); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_335 = x_332; } else { - lean_dec_ref(x_330); - x_333 = lean_box(0); + lean_dec_ref(x_332); + x_335 = lean_box(0); } -if (lean_is_scalar(x_327)) { - x_334 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_329)) { + x_336 = lean_alloc_ctor(0, 2, 0); } else { - x_334 = x_327; + x_336 = x_329; } -lean_ctor_set(x_334, 0, x_326); -lean_ctor_set(x_334, 1, x_331); -if (lean_is_scalar(x_333)) { - x_335 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_336, 0, x_328); +lean_ctor_set(x_336, 1, x_333); +if (lean_is_scalar(x_335)) { + x_337 = lean_alloc_ctor(0, 2, 0); } else { - x_335 = x_333; + x_337 = x_335; } -lean_ctor_set(x_335, 0, x_334); -lean_ctor_set(x_335, 1, x_332); -x_247 = x_335; -goto block_277; +lean_ctor_set(x_337, 0, x_336); +lean_ctor_set(x_337, 1, x_334); +x_249 = x_337; +goto block_279; } else { -lean_object* x_336; lean_object* x_337; lean_object* x_338; lean_object* x_339; -lean_dec(x_327); -lean_dec(x_326); -x_336 = lean_ctor_get(x_330, 0); -lean_inc(x_336); -x_337 = lean_ctor_get(x_330, 1); -lean_inc(x_337); -if (lean_is_exclusive(x_330)) { - lean_ctor_release(x_330, 0); - lean_ctor_release(x_330, 1); - x_338 = x_330; +lean_object* x_338; lean_object* x_339; lean_object* x_340; lean_object* x_341; +lean_dec(x_329); +lean_dec(x_328); +x_338 = lean_ctor_get(x_332, 0); +lean_inc(x_338); +x_339 = lean_ctor_get(x_332, 1); +lean_inc(x_339); +if (lean_is_exclusive(x_332)) { + lean_ctor_release(x_332, 0); + lean_ctor_release(x_332, 1); + x_340 = x_332; } else { - lean_dec_ref(x_330); - x_338 = lean_box(0); + lean_dec_ref(x_332); + x_340 = lean_box(0); } -if (lean_is_scalar(x_338)) { - x_339 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_340)) { + x_341 = lean_alloc_ctor(1, 2, 0); } else { - x_339 = x_338; + x_341 = x_340; } -lean_ctor_set(x_339, 0, x_336); -lean_ctor_set(x_339, 1, x_337); -x_247 = x_339; -goto block_277; +lean_ctor_set(x_341, 0, x_338); +lean_ctor_set(x_341, 1, x_339); +x_249 = x_341; +goto block_279; } } else { -lean_object* x_340; lean_object* x_341; lean_object* x_342; lean_object* x_343; -lean_dec(x_235); +lean_object* x_342; lean_object* x_343; lean_object* x_344; lean_object* x_345; +lean_dec(x_236); lean_dec(x_10); -x_340 = lean_ctor_get(x_322, 0); -lean_inc(x_340); -x_341 = lean_ctor_get(x_322, 1); -lean_inc(x_341); -if (lean_is_exclusive(x_322)) { - lean_ctor_release(x_322, 0); - lean_ctor_release(x_322, 1); - x_342 = x_322; +x_342 = lean_ctor_get(x_324, 0); +lean_inc(x_342); +x_343 = lean_ctor_get(x_324, 1); +lean_inc(x_343); +if (lean_is_exclusive(x_324)) { + lean_ctor_release(x_324, 0); + lean_ctor_release(x_324, 1); + x_344 = x_324; } else { - lean_dec_ref(x_322); - x_342 = lean_box(0); + lean_dec_ref(x_324); + x_344 = lean_box(0); } -if (lean_is_scalar(x_342)) { - x_343 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_344)) { + x_345 = lean_alloc_ctor(1, 2, 0); } else { - x_343 = x_342; + x_345 = x_344; } -lean_ctor_set(x_343, 0, x_340); -lean_ctor_set(x_343, 1, x_341); -x_247 = x_343; -goto block_277; +lean_ctor_set(x_345, 0, x_342); +lean_ctor_set(x_345, 1, x_343); +x_249 = x_345; +goto block_279; } } default: { -lean_object* x_344; lean_object* x_345; lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; -lean_dec(x_294); +lean_object* x_346; lean_object* x_347; lean_object* x_348; lean_object* x_349; lean_object* x_350; lean_object* x_351; lean_object* x_352; lean_object* x_353; lean_object* x_354; lean_object* x_355; +lean_dec(x_296); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_344 = l_Lean_MessageData_ofExpr(x_235); -x_345 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_346 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_346, 0, x_345); -lean_ctor_set(x_346, 1, x_344); -x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_346 = l_Lean_MessageData_ofExpr(x_236); +x_347 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_348 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_348, 0, x_346); -lean_ctor_set(x_348, 1, x_347); -x_349 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_348, 0, x_347); +lean_ctor_set(x_348, 1, x_346); +x_349 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_350 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_350, 0, x_348); lean_ctor_set(x_350, 1, x_349); -x_351 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_351 = l_Lean_MessageData_ofExpr(x_246); x_352 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_352, 0, x_350); lean_ctor_set(x_352, 1, x_351); +x_353 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_354 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_354, 0, x_352); +lean_ctor_set(x_354, 1, x_353); lean_inc(x_12); -x_353 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_352, x_12, x_13, x_14, x_15, x_16, x_17, x_281); -x_247 = x_353; -goto block_277; +x_355 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_354, x_12, x_13, x_14, x_15, x_16, x_17, x_283); +x_249 = x_355; +goto block_279; } } } } 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; -lean_dec(x_246); +lean_object* x_356; lean_object* x_357; lean_object* x_358; lean_object* x_359; lean_object* x_360; lean_object* x_361; lean_object* x_362; lean_object* x_363; lean_object* x_364; lean_object* x_365; +lean_dec(x_248); lean_dec(x_11); lean_dec(x_10); lean_dec(x_9); lean_dec(x_3); lean_dec(x_1); -x_354 = l_Lean_MessageData_ofExpr(x_235); -x_355 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; -x_356 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_356, 0, x_355); -lean_ctor_set(x_356, 1, x_354); -x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; +x_356 = l_Lean_MessageData_ofExpr(x_236); +x_357 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__4___closed__2; x_358 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_358, 0, x_356); -lean_ctor_set(x_358, 1, x_357); -x_359 = l_Lean_MessageData_ofExpr(x_244); +lean_ctor_set(x_358, 0, x_357); +lean_ctor_set(x_358, 1, x_356); +x_359 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__2; x_360 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_360, 0, x_358); lean_ctor_set(x_360, 1, x_359); -x_361 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_361 = l_Lean_MessageData_ofExpr(x_246); x_362 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_362, 0, x_360); lean_ctor_set(x_362, 1, x_361); +x_363 = l_Lean_Elab_Tactic_RCases_rcasesCore___rarg___lambda__8___closed__4; +x_364 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_364, 0, x_362); +lean_ctor_set(x_364, 1, x_363); lean_inc(x_12); -x_363 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_362, x_12, x_13, x_14, x_15, x_16, x_17, x_245); -x_247 = x_363; -goto block_277; +x_365 = l_Lean_throwError___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__3(x_364, x_12, x_13, x_14, x_15, x_16, x_17, x_247); +x_249 = x_365; +goto block_279; } -block_277: +block_279: { -if (lean_obj_tag(x_247) == 0) +if (lean_obj_tag(x_249) == 0) { -lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; uint8_t x_257; -x_248 = lean_ctor_get(x_247, 0); -lean_inc(x_248); -x_249 = lean_ctor_get(x_247, 1); -lean_inc(x_249); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_250 = x_247; -} else { - lean_dec_ref(x_247); - x_250 = lean_box(0); -} -x_251 = lean_ctor_get(x_248, 0); +lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; lean_object* x_257; lean_object* x_258; uint8_t x_259; +x_250 = lean_ctor_get(x_249, 0); +lean_inc(x_250); +x_251 = lean_ctor_get(x_249, 1); lean_inc(x_251); -x_252 = lean_ctor_get(x_248, 1); -lean_inc(x_252); -if (lean_is_exclusive(x_248)) { - lean_ctor_release(x_248, 0); - lean_ctor_release(x_248, 1); - x_253 = x_248; -} else { - lean_dec_ref(x_248); - x_253 = lean_box(0); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_252 = x_249; +} else { + lean_dec_ref(x_249); + x_252 = lean_box(0); +} +x_253 = lean_ctor_get(x_250, 0); +lean_inc(x_253); +x_254 = lean_ctor_get(x_250, 1); +lean_inc(x_254); +if (lean_is_exclusive(x_250)) { + lean_ctor_release(x_250, 0); + lean_ctor_release(x_250, 1); + x_255 = x_250; +} else { + lean_dec_ref(x_250); + x_255 = lean_box(0); } lean_inc(x_6); -if (lean_is_scalar(x_253)) { - x_254 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_255)) { + x_256 = lean_alloc_ctor(0, 2, 0); } else { - x_254 = x_253; + x_256 = x_255; } -lean_ctor_set(x_254, 0, x_251); -lean_ctor_set(x_254, 1, x_6); -x_255 = lean_array_get_size(x_252); -x_256 = lean_unsigned_to_nat(0u); -x_257 = lean_nat_dec_lt(x_256, x_255); -if (x_257 == 0) +lean_ctor_set(x_256, 0, x_253); +lean_ctor_set(x_256, 1, x_6); +x_257 = lean_array_get_size(x_254); +x_258 = lean_unsigned_to_nat(0u); +x_259 = lean_nat_dec_lt(x_258, x_257); +if (x_259 == 0) { -lean_object* x_258; -lean_dec(x_255); +lean_object* x_260; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16382,25 +16391,25 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_258 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_260 = lean_alloc_ctor(0, 2, 0); } else { - x_258 = x_250; + x_260 = x_252; } -lean_ctor_set(x_258, 0, x_6); -lean_ctor_set(x_258, 1, x_249); -return x_258; +lean_ctor_set(x_260, 0, x_6); +lean_ctor_set(x_260, 1, x_251); +return x_260; } else { -uint8_t x_259; -x_259 = lean_nat_dec_le(x_255, x_255); -if (x_259 == 0) +uint8_t x_261; +x_261 = lean_nat_dec_le(x_257, x_257); +if (x_261 == 0) { -lean_object* x_260; -lean_dec(x_255); +lean_object* x_262; +lean_dec(x_257); +lean_dec(x_256); lean_dec(x_254); -lean_dec(x_252); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16410,82 +16419,82 @@ lean_dec(x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); -if (lean_is_scalar(x_250)) { - x_260 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_252)) { + x_262 = lean_alloc_ctor(0, 2, 0); } else { - x_260 = x_250; + x_262 = x_252; } -lean_ctor_set(x_260, 0, x_6); -lean_ctor_set(x_260, 1, x_249); -return x_260; +lean_ctor_set(x_262, 0, x_6); +lean_ctor_set(x_262, 1, x_251); +return x_262; } else { -size_t x_261; size_t x_262; lean_object* x_263; -lean_dec(x_250); -lean_dec(x_6); -x_261 = 0; -x_262 = lean_usize_of_nat(x_255); -lean_dec(x_255); -x_263 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_252, x_261, x_262, x_254, x_12, x_13, x_14, x_15, x_16, x_17, x_249); +size_t x_263; size_t x_264; lean_object* x_265; lean_dec(x_252); -if (lean_obj_tag(x_263) == 0) -{ -lean_object* x_264; lean_object* x_265; lean_object* x_266; lean_object* x_267; lean_object* x_268; -x_264 = lean_ctor_get(x_263, 0); -lean_inc(x_264); -x_265 = lean_ctor_get(x_263, 1); -lean_inc(x_265); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_266 = x_263; -} else { - lean_dec_ref(x_263); - x_266 = lean_box(0); -} -x_267 = lean_ctor_get(x_264, 1); +lean_dec(x_6); +x_263 = 0; +x_264 = lean_usize_of_nat(x_257); +lean_dec(x_257); +x_265 = l_Array_foldlMUnsafe_fold___at_Lean_Elab_Tactic_RCases_rcasesCore___spec__8___rarg(x_4, x_7, x_8, x_254, x_263, x_264, x_256, x_12, x_13, x_14, x_15, x_16, x_17, x_251); +lean_dec(x_254); +if (lean_obj_tag(x_265) == 0) +{ +lean_object* x_266; lean_object* x_267; lean_object* x_268; lean_object* x_269; lean_object* x_270; +x_266 = lean_ctor_get(x_265, 0); +lean_inc(x_266); +x_267 = lean_ctor_get(x_265, 1); lean_inc(x_267); -lean_dec(x_264); -if (lean_is_scalar(x_266)) { - x_268 = lean_alloc_ctor(0, 2, 0); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_268 = x_265; +} else { + lean_dec_ref(x_265); + x_268 = lean_box(0); +} +x_269 = lean_ctor_get(x_266, 1); +lean_inc(x_269); +lean_dec(x_266); +if (lean_is_scalar(x_268)) { + x_270 = lean_alloc_ctor(0, 2, 0); } else { - x_268 = x_266; + x_270 = x_268; } -lean_ctor_set(x_268, 0, x_267); -lean_ctor_set(x_268, 1, x_265); -return x_268; +lean_ctor_set(x_270, 0, x_269); +lean_ctor_set(x_270, 1, x_267); +return x_270; } else { -lean_object* x_269; lean_object* x_270; lean_object* x_271; lean_object* x_272; -x_269 = lean_ctor_get(x_263, 0); -lean_inc(x_269); -x_270 = lean_ctor_get(x_263, 1); -lean_inc(x_270); -if (lean_is_exclusive(x_263)) { - lean_ctor_release(x_263, 0); - lean_ctor_release(x_263, 1); - x_271 = x_263; +lean_object* x_271; lean_object* x_272; lean_object* x_273; lean_object* x_274; +x_271 = lean_ctor_get(x_265, 0); +lean_inc(x_271); +x_272 = lean_ctor_get(x_265, 1); +lean_inc(x_272); +if (lean_is_exclusive(x_265)) { + lean_ctor_release(x_265, 0); + lean_ctor_release(x_265, 1); + x_273 = x_265; } else { - lean_dec_ref(x_263); - x_271 = lean_box(0); + lean_dec_ref(x_265); + x_273 = lean_box(0); } -if (lean_is_scalar(x_271)) { - x_272 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_273)) { + x_274 = lean_alloc_ctor(1, 2, 0); } else { - x_272 = x_271; + x_274 = x_273; } -lean_ctor_set(x_272, 0, x_269); -lean_ctor_set(x_272, 1, x_270); -return x_272; +lean_ctor_set(x_274, 0, x_271); +lean_ctor_set(x_274, 1, x_272); +return x_274; } } } } else { -lean_object* x_273; lean_object* x_274; lean_object* x_275; lean_object* x_276; +lean_object* x_275; lean_object* x_276; lean_object* x_277; lean_object* x_278; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16496,33 +16505,33 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_4); -x_273 = lean_ctor_get(x_247, 0); -lean_inc(x_273); -x_274 = lean_ctor_get(x_247, 1); -lean_inc(x_274); -if (lean_is_exclusive(x_247)) { - lean_ctor_release(x_247, 0); - lean_ctor_release(x_247, 1); - x_275 = x_247; +x_275 = lean_ctor_get(x_249, 0); +lean_inc(x_275); +x_276 = lean_ctor_get(x_249, 1); +lean_inc(x_276); +if (lean_is_exclusive(x_249)) { + lean_ctor_release(x_249, 0); + lean_ctor_release(x_249, 1); + x_277 = x_249; } else { - lean_dec_ref(x_247); - x_275 = lean_box(0); + lean_dec_ref(x_249); + x_277 = lean_box(0); } -if (lean_is_scalar(x_275)) { - x_276 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_277)) { + x_278 = lean_alloc_ctor(1, 2, 0); } else { - x_276 = x_275; + x_278 = x_277; } -lean_ctor_set(x_276, 0, x_273); -lean_ctor_set(x_276, 1, x_274); -return x_276; +lean_ctor_set(x_278, 0, x_275); +lean_ctor_set(x_278, 1, x_276); +return x_278; } } } else { -lean_object* x_364; lean_object* x_365; lean_object* x_366; lean_object* x_367; -lean_dec(x_235); +lean_object* x_366; lean_object* x_367; lean_object* x_368; lean_object* x_369; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16538,32 +16547,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_364 = lean_ctor_get(x_243, 0); -lean_inc(x_364); -x_365 = lean_ctor_get(x_243, 1); -lean_inc(x_365); -if (lean_is_exclusive(x_243)) { - lean_ctor_release(x_243, 0); - lean_ctor_release(x_243, 1); - x_366 = x_243; +x_366 = lean_ctor_get(x_245, 0); +lean_inc(x_366); +x_367 = lean_ctor_get(x_245, 1); +lean_inc(x_367); +if (lean_is_exclusive(x_245)) { + lean_ctor_release(x_245, 0); + lean_ctor_release(x_245, 1); + x_368 = x_245; } else { - lean_dec_ref(x_243); - x_366 = lean_box(0); + lean_dec_ref(x_245); + x_368 = lean_box(0); } -if (lean_is_scalar(x_366)) { - x_367 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_368)) { + x_369 = lean_alloc_ctor(1, 2, 0); } else { - x_367 = x_366; + x_369 = x_368; } -lean_ctor_set(x_367, 0, x_364); -lean_ctor_set(x_367, 1, x_365); -return x_367; +lean_ctor_set(x_369, 0, x_366); +lean_ctor_set(x_369, 1, x_367); +return x_369; } } else { -lean_object* x_368; lean_object* x_369; lean_object* x_370; lean_object* x_371; -lean_dec(x_235); +lean_object* x_370; lean_object* x_371; lean_object* x_372; lean_object* x_373; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16579,32 +16588,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_368 = lean_ctor_get(x_240, 0); -lean_inc(x_368); -x_369 = lean_ctor_get(x_240, 1); -lean_inc(x_369); -if (lean_is_exclusive(x_240)) { - lean_ctor_release(x_240, 0); - lean_ctor_release(x_240, 1); - x_370 = x_240; +x_370 = lean_ctor_get(x_242, 0); +lean_inc(x_370); +x_371 = lean_ctor_get(x_242, 1); +lean_inc(x_371); +if (lean_is_exclusive(x_242)) { + lean_ctor_release(x_242, 0); + lean_ctor_release(x_242, 1); + x_372 = x_242; } else { - lean_dec_ref(x_240); - x_370 = lean_box(0); + lean_dec_ref(x_242); + x_372 = lean_box(0); } -if (lean_is_scalar(x_370)) { - x_371 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_372)) { + x_373 = lean_alloc_ctor(1, 2, 0); } else { - x_371 = x_370; + x_373 = x_372; } -lean_ctor_set(x_371, 0, x_368); -lean_ctor_set(x_371, 1, x_369); -return x_371; +lean_ctor_set(x_373, 0, x_370); +lean_ctor_set(x_373, 1, x_371); +return x_373; } } else { -lean_object* x_372; lean_object* x_373; lean_object* x_374; lean_object* x_375; -lean_dec(x_235); +lean_object* x_374; lean_object* x_375; lean_object* x_376; lean_object* x_377; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16620,32 +16629,32 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_372 = lean_ctor_get(x_238, 0); -lean_inc(x_372); -x_373 = lean_ctor_get(x_238, 1); -lean_inc(x_373); -if (lean_is_exclusive(x_238)) { - lean_ctor_release(x_238, 0); - lean_ctor_release(x_238, 1); - x_374 = x_238; +x_374 = lean_ctor_get(x_240, 0); +lean_inc(x_374); +x_375 = lean_ctor_get(x_240, 1); +lean_inc(x_375); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + x_376 = x_240; } else { - lean_dec_ref(x_238); - x_374 = lean_box(0); + lean_dec_ref(x_240); + x_376 = lean_box(0); } -if (lean_is_scalar(x_374)) { - x_375 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_376)) { + x_377 = lean_alloc_ctor(1, 2, 0); } else { - x_375 = x_374; + x_377 = x_376; } -lean_ctor_set(x_375, 0, x_372); -lean_ctor_set(x_375, 1, x_373); -return x_375; +lean_ctor_set(x_377, 0, x_374); +lean_ctor_set(x_377, 1, x_375); +return x_377; } } else { -lean_object* x_376; lean_object* x_377; lean_object* x_378; lean_object* x_379; -lean_dec(x_235); +lean_object* x_378; lean_object* x_379; lean_object* x_380; lean_object* x_381; +lean_dec(x_236); lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16661,33 +16670,33 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_376 = lean_ctor_get(x_236, 0); -lean_inc(x_376); -x_377 = lean_ctor_get(x_236, 1); -lean_inc(x_377); -if (lean_is_exclusive(x_236)) { - lean_ctor_release(x_236, 0); - lean_ctor_release(x_236, 1); - x_378 = x_236; +x_378 = lean_ctor_get(x_237, 0); +lean_inc(x_378); +x_379 = lean_ctor_get(x_237, 1); +lean_inc(x_379); +if (lean_is_exclusive(x_237)) { + lean_ctor_release(x_237, 0); + lean_ctor_release(x_237, 1); + x_380 = x_237; } else { - lean_dec_ref(x_236); - x_378 = lean_box(0); + lean_dec_ref(x_237); + x_380 = lean_box(0); } -if (lean_is_scalar(x_378)) { - x_379 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_380)) { + x_381 = lean_alloc_ctor(1, 2, 0); } else { - x_379 = x_378; + x_381 = x_380; } -lean_ctor_set(x_379, 0, x_376); -lean_ctor_set(x_379, 1, x_377); -return x_379; +lean_ctor_set(x_381, 0, x_378); +lean_ctor_set(x_381, 1, x_379); +return x_381; } } } } else { -uint8_t x_389; +uint8_t x_391; lean_dec(x_17); lean_dec(x_16); lean_dec(x_15); @@ -16704,23 +16713,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_389 = !lean_is_exclusive(x_21); -if (x_389 == 0) +x_391 = !lean_is_exclusive(x_21); +if (x_391 == 0) { return x_21; } else { -lean_object* x_390; lean_object* x_391; lean_object* x_392; -x_390 = lean_ctor_get(x_21, 0); -x_391 = lean_ctor_get(x_21, 1); -lean_inc(x_391); -lean_inc(x_390); +lean_object* x_392; lean_object* x_393; lean_object* x_394; +x_392 = lean_ctor_get(x_21, 0); +x_393 = lean_ctor_get(x_21, 1); +lean_inc(x_393); +lean_inc(x_392); lean_dec(x_21); -x_392 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_392, 0, x_390); -lean_ctor_set(x_392, 1, x_391); -return x_392; +x_394 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_394, 0, x_392); +lean_ctor_set(x_394, 1, x_393); +return x_394; } } } @@ -27652,12 +27661,12 @@ lean_closure_set(x_25, 0, x_1); lean_closure_set(x_25, 1, x_3); lean_closure_set(x_25, 2, x_24); x_26 = 1; -x_27 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_25, x_22, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_27 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_25, x_26, x_4, x_5, x_6, x_7, x_8, x_9, x_10); return x_27; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; uint8_t x_32; lean_object* x_33; +lean_object* x_28; lean_object* x_29; lean_object* x_30; uint8_t x_31; lean_object* x_32; lean_dec(x_11); x_28 = lean_box(0); x_29 = lean_alloc_ctor(1, 2, 0); @@ -27667,28 +27676,26 @@ x_30 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_RCases_rcases___lambda__1__ lean_closure_set(x_30, 0, x_1); lean_closure_set(x_30, 1, x_3); lean_closure_set(x_30, 2, x_29); -x_31 = 0; -x_32 = 1; -x_33 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_30, x_31, x_32, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_33; +x_31 = 1; +x_32 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_30, x_31, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_32; } } else { -lean_object* x_34; lean_object* x_35; lean_object* x_36; uint8_t x_37; uint8_t x_38; lean_object* x_39; +lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; lean_object* x_37; lean_dec(x_11); lean_dec(x_2); lean_dec(x_1); -x_34 = lean_box(0); -x_35 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_35, 0, x_3); -lean_ctor_set(x_35, 1, x_34); -x_36 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed), 8, 1); -lean_closure_set(x_36, 0, x_35); -x_37 = 0; -x_38 = 1; -x_39 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_36, x_37, x_38, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_39; +x_33 = lean_box(0); +x_34 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_34, 0, x_3); +lean_ctor_set(x_34, 1, x_33); +x_35 = lean_alloc_closure((void*)(l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg___boxed), 8, 1); +lean_closure_set(x_35, 0, x_34); +x_36 = 1; +x_37 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_35, x_36, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_37; } } } @@ -28001,15 +28008,14 @@ return x_64; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_obtainNone(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, 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; lean_object* x_14; +lean_object* x_11; uint8_t x_12; lean_object* x_13; x_11 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_RCases_obtainNone___lambda__1), 10, 3); lean_closure_set(x_11, 0, x_2); lean_closure_set(x_11, 1, x_3); lean_closure_set(x_11, 2, x_1); -x_12 = 0; -x_13 = 1; -x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_11, x_12, x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_14; +x_12 = 1; +x_13 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_11, x_12, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_13; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_expandRIntroPat___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { @@ -28864,7 +28870,7 @@ return x_26; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_RCases_rintro(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_t x_16; lean_object* x_17; +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; x_11 = lean_box(0); x_12 = l_Lean_Elab_Tactic_RCases_instCoeTSyntaxConsSyntaxNodeKindMkStr4Nil__lean___closed__3; x_13 = l_Lean_Elab_Tactic_RCases_obtainNone___lambda__1___closed__1; @@ -28875,10 +28881,9 @@ lean_closure_set(x_14, 2, x_12); lean_closure_set(x_14, 3, x_1); lean_closure_set(x_14, 4, x_2); lean_closure_set(x_14, 5, x_13); -x_15 = 0; -x_16 = 1; -x_17 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_14, x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -return x_17; +x_15 = 1; +x_16 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_14, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +return x_16; } } LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_RCases_evalRCases___spec__1(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c b/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c index d9e2df12e671..fdc6895e17ab 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Rewrite.c @@ -79,7 +79,7 @@ static lean_object* l_Lean_Elab_Tactic_elabRewriteConfig___closed__6; static lean_object* l_Lean_Elab_Tactic_elabRewriteConfig___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_1164____boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq_declRange___closed__3; -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_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*); size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_Elab_Tactic_evalRewriteSeq___lambda__3___closed__1; 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*); @@ -128,7 +128,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_evalRewriteSeq___lambda__1___boxed(l LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl___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_Tactic_evalRewriteSeq___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*); static lean_object* l_Std_Range_forIn_loop___at_Lean_Elab_Tactic_withRWRulesSeq___spec__1___lambda__4___closed__2; -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_elabRewriteConfig___closed__4; static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalRewriteSeq___closed__2; @@ -516,7 +516,7 @@ return x_61; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteTarget(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10, lean_object* x_11, lean_object* x_12) { _start: { -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; uint8_t x_18; lean_object* x_19; +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; x_13 = lean_box(0); x_14 = lean_box(x_2); x_15 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_rewriteTarget___lambda__1___boxed), 13, 4); @@ -528,10 +528,9 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_withMainContext___rarg), 10 lean_closure_set(x_16, 0, x_15); lean_closure_set(x_16, 1, x_4); lean_closure_set(x_16, 2, x_5); -x_17 = 0; -x_18 = 1; -x_19 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_16, x_17, x_18, x_6, x_7, x_8, x_9, x_10, x_11, x_12); -return x_19; +x_17 = 1; +x_18 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_16, x_17, x_6, x_7, x_8, x_9, x_10, x_11, x_12); +return x_18; } } LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteTarget___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) { @@ -704,30 +703,29 @@ return x_38; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_rewriteLocalDecl___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; uint8_t x_14; lean_object* x_15; +lean_object* x_12; uint8_t x_13; lean_object* x_14; lean_inc(x_4); lean_inc(x_3); x_12 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_withMainContext___rarg), 10, 3); lean_closure_set(x_12, 0, x_1); lean_closure_set(x_12, 1, x_3); lean_closure_set(x_12, 2, x_4); -x_13 = 0; -x_14 = 1; +x_13 = 1; lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); -x_15 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_14, x_5, x_6, x_7, x_8, x_9, x_10, x_11); -if (lean_obj_tag(x_15) == 0) +x_14 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_12, x_13, x_5, x_6, x_7, x_8, x_9, x_10, x_11); +if (lean_obj_tag(x_14) == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; -x_16 = lean_ctor_get(x_15, 0); +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_14, 0); +lean_inc(x_15); +x_16 = lean_ctor_get(x_14, 1); lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); -lean_inc(x_17); -lean_dec(x_15); +lean_dec(x_14); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); @@ -736,48 +734,48 @@ lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_18 = l_Lean_Elab_Tactic_getMainGoal(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_17); -if (lean_obj_tag(x_18) == 0) +x_17 = l_Lean_Elab_Tactic_getMainGoal(x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_16); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_19 = lean_ctor_get(x_18, 0); +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_17, 0); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 1); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 1); +lean_dec(x_17); +x_20 = lean_ctor_get(x_15, 0); lean_inc(x_20); -lean_dec(x_18); -x_21 = lean_ctor_get(x_16, 0); +x_21 = lean_ctor_get(x_15, 1); lean_inc(x_21); -x_22 = lean_ctor_get(x_16, 1); -lean_inc(x_22); lean_inc(x_10); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); -x_23 = l___private_Lean_Meta_Tactic_Replace_0__Lean_Meta_replaceLocalDeclCore(x_19, x_2, x_21, x_22, x_7, x_8, x_9, x_10, x_20); -if (lean_obj_tag(x_23) == 0) +x_22 = l___private_Lean_Meta_Tactic_Replace_0__Lean_Meta_replaceLocalDeclCore(x_18, x_2, x_20, x_21, x_7, x_8, x_9, x_10, x_19); +if (lean_obj_tag(x_22) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_24 = lean_ctor_get(x_23, 0); +lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; +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, 1); lean_inc(x_25); lean_dec(x_23); -x_26 = lean_ctor_get(x_24, 1); +x_26 = lean_ctor_get(x_15, 2); lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_ctor_get(x_16, 2); -lean_inc(x_27); -lean_dec(x_16); -x_28 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_28, 0, x_26); -lean_ctor_set(x_28, 1, x_27); -x_29 = l_Lean_Elab_Tactic_replaceMainGoal(x_28, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_25); -return x_29; +lean_dec(x_15); +x_27 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_27, 0, x_25); +lean_ctor_set(x_27, 1, x_26); +x_28 = l_Lean_Elab_Tactic_replaceMainGoal(x_27, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10, x_24); +return x_28; } else { -uint8_t x_30; -lean_dec(x_16); +uint8_t x_29; +lean_dec(x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -786,30 +784,30 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_30 = !lean_is_exclusive(x_23); -if (x_30 == 0) +x_29 = !lean_is_exclusive(x_22); +if (x_29 == 0) { -return x_23; +return x_22; } 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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_22, 0); +x_31 = lean_ctor_get(x_22, 1); lean_inc(x_31); -lean_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_inc(x_30); +lean_dec(x_22); +x_32 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_31); +return x_32; } } } else { -uint8_t x_34; -lean_dec(x_16); +uint8_t x_33; +lean_dec(x_15); lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -819,29 +817,29 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = !lean_is_exclusive(x_18); -if (x_34 == 0) +x_33 = !lean_is_exclusive(x_17); +if (x_33 == 0) { -return x_18; +return x_17; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; -x_35 = lean_ctor_get(x_18, 0); -x_36 = lean_ctor_get(x_18, 1); -lean_inc(x_36); +lean_object* x_34; lean_object* x_35; lean_object* x_36; +x_34 = lean_ctor_get(x_17, 0); +x_35 = lean_ctor_get(x_17, 1); lean_inc(x_35); -lean_dec(x_18); -x_37 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_37, 0, x_35); -lean_ctor_set(x_37, 1, x_36); -return x_37; +lean_inc(x_34); +lean_dec(x_17); +x_36 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_36, 0, x_34); +lean_ctor_set(x_36, 1, x_35); +return x_36; } } } else { -uint8_t x_38; +uint8_t x_37; lean_dec(x_10); lean_dec(x_9); lean_dec(x_8); @@ -851,23 +849,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_38 = !lean_is_exclusive(x_15); -if (x_38 == 0) +x_37 = !lean_is_exclusive(x_14); +if (x_37 == 0) { -return x_15; +return x_14; } else { -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_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_14, 0); +x_39 = lean_ctor_get(x_14, 1); 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; +lean_inc(x_38); +lean_dec(x_14); +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; } } } @@ -7122,38 +7120,39 @@ lean_inc(x_4); x_12 = l_Lean_Elab_Term_elabTermEnsuringType(x_1, x_2, x_11, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_18; +return x_19; } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -7161,52 +7160,52 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -return x_16; +return x_17; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -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_17); +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_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_12); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) { return x_12; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -x_25 = lean_ctor_get(x_12, 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_12); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_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; } } } @@ -7339,7 +7338,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -7351,85 +7350,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabRewriteConfig___lambda_ lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabRewriteConfig___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__7; -x_24 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabRewriteConfig___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__7; +x_22 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabRewriteConfig___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_1164_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_evalUnsafe____x40_Lean_Elab_Tactic_Rewrite___hyg_1164_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__11; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_elabRewriteConfig___closed__11; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Rewrites.c b/stage0/stdlib/Lean/Elab/Tactic/Rewrites.c index 4e870dcc4c2f..eab1fdf06109 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Rewrites.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Rewrites.c @@ -1043,6 +1043,7 @@ lean_inc(x_33); lean_dec(x_31); x_34 = l_Lean_Elab_Rewrites_evalExact___lambda__2___closed__2; x_35 = lean_unsigned_to_nat(90u); +lean_inc(x_5); x_36 = l_Lean_reportOutOfHeartbeats(x_34, x_5, x_35, x_13, x_14, x_33); x_37 = lean_ctor_get(x_36, 1); lean_inc(x_37); @@ -2221,6 +2222,7 @@ lean_inc(x_30); lean_dec(x_28); x_31 = l_Lean_Elab_Rewrites_evalExact___lambda__2___closed__2; x_32 = lean_unsigned_to_nat(90u); +lean_inc(x_4); x_33 = l_Lean_reportOutOfHeartbeats(x_31, x_4, x_32, x_11, x_12, x_30); x_34 = lean_ctor_get(x_33, 1); lean_inc(x_34); @@ -2483,6 +2485,7 @@ lean_dec(x_16); x_19 = 0; x_20 = l_Lean_Elab_Rewrites_evalExact___lambda__7___closed__3; x_21 = lean_unsigned_to_nat(90u); +lean_inc(x_1); x_22 = l_Lean_reportOutOfHeartbeats(x_20, x_1, x_21, x_11, x_12, x_18); if (lean_obj_tag(x_4) == 0) { diff --git a/stage0/stdlib/Lean/Elab/Tactic/ShowTerm.c b/stage0/stdlib/Lean/Elab/Tactic/ShowTerm.c index 1056b4529f1d..acaf5153f73c 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/ShowTerm.c +++ b/stage0/stdlib/Lean/Elab/Tactic/ShowTerm.c @@ -586,93 +586,94 @@ lean_inc(x_3); x_19 = l_Lean_Elab_Term_elabTermEnsuringType(x_16, x_2, x_18, x_18, x_17, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_19) == 0) { -lean_object* x_20; lean_object* x_21; uint8_t x_22; lean_object* x_23; +lean_object* x_20; lean_object* x_21; uint8_t x_22; uint8_t 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 = 0; +x_22 = 1; +x_23 = 0; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_23 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_22, x_22, x_3, x_4, x_5, x_6, x_7, x_8, x_21); -if (lean_obj_tag(x_23) == 0) +x_24 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_22, x_23, x_3, x_4, x_5, x_6, x_7, x_8, x_21); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; -x_24 = lean_ctor_get(x_23, 1); -lean_inc(x_24); -lean_dec(x_23); +lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; +x_25 = lean_ctor_get(x_24, 1); +lean_inc(x_25); +lean_dec(x_24); lean_inc(x_20); -x_25 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_24); +x_26 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_20, x_3, x_4, x_5, x_6, x_7, x_8, x_25); lean_dec(x_4); lean_dec(x_3); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -x_28 = lean_ctor_get(x_7, 5); +x_28 = lean_ctor_get(x_26, 1); lean_inc(x_28); -x_29 = l_Lean_Expr_headBeta(x_26); -x_30 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_30, 0, x_28); -x_31 = l_Lean_Elab_Tactic_ShowTerm_elabShowTerm___closed__4; -x_32 = l_Lean_Meta_Tactic_TryThis_addTermSuggestion(x_14, x_29, x_30, x_31, x_17, x_5, x_6, x_7, x_8, x_27); -if (lean_obj_tag(x_32) == 0) -{ -uint8_t x_33; -x_33 = !lean_is_exclusive(x_32); -if (x_33 == 0) -{ -lean_object* x_34; -x_34 = lean_ctor_get(x_32, 0); -lean_dec(x_34); -lean_ctor_set(x_32, 0, x_20); -return x_32; +lean_dec(x_26); +x_29 = lean_ctor_get(x_7, 5); +lean_inc(x_29); +x_30 = l_Lean_Expr_headBeta(x_27); +x_31 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_31, 0, x_29); +x_32 = l_Lean_Elab_Tactic_ShowTerm_elabShowTerm___closed__4; +x_33 = l_Lean_Meta_Tactic_TryThis_addTermSuggestion(x_14, x_30, x_31, x_32, x_17, x_5, x_6, x_7, x_8, x_28); +if (lean_obj_tag(x_33) == 0) +{ +uint8_t x_34; +x_34 = !lean_is_exclusive(x_33); +if (x_34 == 0) +{ +lean_object* x_35; +x_35 = lean_ctor_get(x_33, 0); +lean_dec(x_35); +lean_ctor_set(x_33, 0, x_20); +return x_33; } else { -lean_object* x_35; lean_object* x_36; -x_35 = lean_ctor_get(x_32, 1); -lean_inc(x_35); -lean_dec(x_32); -x_36 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_20); -lean_ctor_set(x_36, 1, x_35); -return x_36; +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_20); +lean_ctor_set(x_37, 1, x_36); +return x_37; } } else { -uint8_t x_37; +uint8_t x_38; lean_dec(x_20); -x_37 = !lean_is_exclusive(x_32); -if (x_37 == 0) +x_38 = !lean_is_exclusive(x_33); +if (x_38 == 0) { -return x_32; +return x_33; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_38 = lean_ctor_get(x_32, 0); -x_39 = lean_ctor_get(x_32, 1); +lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_39 = lean_ctor_get(x_33, 0); +x_40 = lean_ctor_get(x_33, 1); +lean_inc(x_40); lean_inc(x_39); -lean_inc(x_38); -lean_dec(x_32); -x_40 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_40, 0, x_38); -lean_ctor_set(x_40, 1, x_39); -return x_40; +lean_dec(x_33); +x_41 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_41, 0, x_39); +lean_ctor_set(x_41, 1, x_40); +return x_41; } } } else { -uint8_t x_41; +uint8_t x_42; lean_dec(x_20); lean_dec(x_14); lean_dec(x_8); @@ -681,29 +682,29 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_23); -if (x_41 == 0) +x_42 = !lean_is_exclusive(x_24); +if (x_42 == 0) { -return x_23; +return x_24; } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_42 = lean_ctor_get(x_23, 0); -x_43 = lean_ctor_get(x_23, 1); +lean_object* x_43; lean_object* x_44; lean_object* x_45; +x_43 = lean_ctor_get(x_24, 0); +x_44 = lean_ctor_get(x_24, 1); +lean_inc(x_44); lean_inc(x_43); -lean_inc(x_42); -lean_dec(x_23); -x_44 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_44, 0, x_42); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_dec(x_24); +x_45 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_45, 0, x_43); +lean_ctor_set(x_45, 1, x_44); +return x_45; } } } else { -uint8_t x_45; +uint8_t x_46; lean_dec(x_14); lean_dec(x_8); lean_dec(x_7); @@ -711,23 +712,23 @@ lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_45 = !lean_is_exclusive(x_19); -if (x_45 == 0) +x_46 = !lean_is_exclusive(x_19); +if (x_46 == 0) { return x_19; } else { -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_object* x_47; lean_object* x_48; lean_object* x_49; +x_47 = lean_ctor_get(x_19, 0); +x_48 = lean_ctor_get(x_19, 1); +lean_inc(x_48); 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; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/Simp.c b/stage0/stdlib/Lean/Elab/Tactic/Simp.c index 375808d98d29..40216357366d 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Simp.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Simp.c @@ -185,7 +185,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_elabSimpCo LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SimpKind_noConfusion(lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_Elab_Tactic_elabSimpArgs___spec__3___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* 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_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_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_Elab_Tactic_elabSimpArgs___spec__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 l_Lean_Expr_hasMVar(lean_object*); @@ -352,7 +352,7 @@ static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpOnly___ LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Tactic_mkSimpContext___lambda__4___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_mkSimpContext___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Elab_Tactic_evalSimp_declRange___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Tactic_mkSimpOnly___spec__9___closed__8; static lean_object* l_Lean_Parser_Tactic_getDSimpArgs_x3f___closed__1; @@ -2924,38 +2924,39 @@ lean_inc(x_4); x_12 = l_Lean_Elab_Term_elabTermEnsuringType(x_1, x_2, x_11, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_18; +return x_19; } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -2963,52 +2964,52 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -return x_16; +return x_17; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -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_17); +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_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_12); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) { return x_12; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -x_25 = lean_ctor_get(x_12, 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_12); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_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; } } } @@ -3159,7 +3160,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -3171,85 +3172,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabSimpConfigCore___lambda lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; -x_24 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; +x_22 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_6_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__11; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__11; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } } } @@ -3456,7 +3454,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -3468,85 +3466,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabSimpConfigCore___lambda lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; -x_24 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; +x_22 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_191_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_191_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_elabSimpConfigCtxCore___closed__3; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_elabSimpConfigCtxCore___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_8); +return x_33; } } } @@ -3648,7 +3643,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -3660,85 +3655,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_elabSimpConfigCore___lambda lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; -x_24 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__7; +x_22 = l_Lean_Elab_Tactic_elabSimpConfigCore___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_elabSimpConfigCore___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_376_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_evalUnsafe____x40_Lean_Elab_Tactic_Simp___hyg_376_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_elabDSimpConfigCore___closed__3; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_elabDSimpConfigCore___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_8); +return x_33; } } } @@ -3941,86 +3933,85 @@ return x_17; LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_tacticToDischarge___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { _start: { -uint8_t x_10; uint8_t x_11; lean_object* x_12; -x_10 = 0; -x_11 = 1; +uint8_t x_10; lean_object* x_11; +x_10 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); -x_12 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_1, x_10, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_12) == 0) +x_11 = l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg(x_1, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +if (lean_obj_tag(x_11) == 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 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +lean_object* x_12; lean_object* x_13; uint8_t x_14; +x_12 = lean_ctor_get(x_11, 1); +lean_inc(x_12); +lean_dec(x_11); +x_13 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_12); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +x_14 = !lean_is_exclusive(x_13); +if (x_14 == 0) { -lean_object* x_16; uint8_t x_17; -x_16 = lean_ctor_get(x_14, 0); -x_17 = l_Lean_Expr_hasExprMVar(x_16); -if (x_17 == 0) +lean_object* x_15; uint8_t x_16; +x_15 = lean_ctor_get(x_13, 0); +x_16 = l_Lean_Expr_hasExprMVar(x_15); +if (x_16 == 0) { -lean_object* x_18; -x_18 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_18, 0, x_16); -lean_ctor_set(x_14, 0, x_18); -return x_14; +lean_object* x_17; +x_17 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_17, 0, x_15); +lean_ctor_set(x_13, 0, x_17); +return x_13; } else { -lean_object* x_19; -lean_dec(x_16); -x_19 = lean_box(0); -lean_ctor_set(x_14, 0, x_19); -return x_14; +lean_object* x_18; +lean_dec(x_15); +x_18 = lean_box(0); +lean_ctor_set(x_13, 0, x_18); +return x_13; } } else { -lean_object* x_20; lean_object* x_21; uint8_t x_22; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); +lean_object* x_19; lean_object* x_20; uint8_t x_21; +x_19 = lean_ctor_get(x_13, 0); +x_20 = lean_ctor_get(x_13, 1); lean_inc(x_20); -lean_dec(x_14); -x_22 = l_Lean_Expr_hasExprMVar(x_20); -if (x_22 == 0) +lean_inc(x_19); +lean_dec(x_13); +x_21 = l_Lean_Expr_hasExprMVar(x_19); +if (x_21 == 0) { -lean_object* x_23; lean_object* x_24; -x_23 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_23, 0, x_20); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_21); -return x_24; +lean_object* x_22; lean_object* x_23; +x_22 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_22, 0, x_19); +x_23 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_23, 0, x_22); +lean_ctor_set(x_23, 1, x_20); +return x_23; } else { -lean_object* x_25; lean_object* x_26; -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_21); -return x_26; +lean_object* x_24; lean_object* x_25; +lean_dec(x_19); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_20); +return x_25; } } } else { -uint8_t x_27; +uint8_t x_26; lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); @@ -4028,23 +4019,23 @@ lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = !lean_is_exclusive(x_12); -if (x_27 == 0) +x_26 = !lean_is_exclusive(x_11); +if (x_26 == 0) { -return x_12; +return x_11; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_12, 0); -x_29 = lean_ctor_get(x_12, 1); -lean_inc(x_29); +lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_27 = lean_ctor_get(x_11, 0); +x_28 = lean_ctor_get(x_11, 1); lean_inc(x_28); -lean_dec(x_12); -x_30 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_30, 0, x_28); -lean_ctor_set(x_30, 1, x_29); -return x_30; +lean_inc(x_27); +lean_dec(x_11); +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; } } } @@ -5540,7 +5531,7 @@ lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); -x_14 = 0; +x_14 = 1; lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); diff --git a/stage0/stdlib/Lean/Elab/Tactic/Simproc.c b/stage0/stdlib/Lean/Elab/Tactic/Simproc.c index 9e18fddce729..99ad7f1bdcf3 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/Simproc.c +++ b/stage0/stdlib/Lean/Elab/Tactic/Simproc.c @@ -196,88 +196,89 @@ lean_inc(x_3); x_11 = l_Lean_Elab_Term_elabTerm(x_1, x_2, x_10, x_10, x_3, x_4, x_5, x_6, x_7, x_8, x_9); if (lean_obj_tag(x_11) == 0) { -lean_object* x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; +lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; x_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); lean_dec(x_11); x_14 = 0; -x_15 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_10, x_14, x_3, x_4, x_5, x_6, x_7, x_8, x_13); -if (lean_obj_tag(x_15) == 0) +x_15 = 0; +x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_14, x_15, x_3, x_4, x_5, x_6, x_7, x_8, x_13); +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; -x_17 = lean_ctor_get(x_15, 0); -lean_dec(x_17); -lean_ctor_set(x_15, 0, x_12); -return x_15; +lean_object* x_18; +x_18 = lean_ctor_get(x_16, 0); +lean_dec(x_18); +lean_ctor_set(x_16, 0, x_12); +return x_16; } else { -lean_object* x_18; lean_object* x_19; -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_19, 0, x_12); -lean_ctor_set(x_19, 1, x_18); -return x_19; +lean_object* x_19; lean_object* x_20; +x_19 = lean_ctor_get(x_16, 1); +lean_inc(x_19); +lean_dec(x_16); +x_20 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_20, 0, x_12); +lean_ctor_set(x_20, 1, x_19); +return x_20; } } else { -uint8_t x_20; +uint8_t x_21; lean_dec(x_12); -x_20 = !lean_is_exclusive(x_15); -if (x_20 == 0) +x_21 = !lean_is_exclusive(x_16); +if (x_21 == 0) { -return x_15; +return x_16; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_15, 0); -x_22 = lean_ctor_get(x_15, 1); +lean_object* x_22; lean_object* x_23; lean_object* x_24; +x_22 = lean_ctor_get(x_16, 0); +x_23 = lean_ctor_get(x_16, 1); +lean_inc(x_23); lean_inc(x_22); -lean_inc(x_21); -lean_dec(x_15); -x_23 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_23, 0, x_21); -lean_ctor_set(x_23, 1, x_22); -return x_23; +lean_dec(x_16); +x_24 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_24, 0, x_22); +lean_ctor_set(x_24, 1, x_23); +return x_24; } } } else { -uint8_t x_24; +uint8_t x_25; 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_24 = !lean_is_exclusive(x_11); -if (x_24 == 0) +x_25 = !lean_is_exclusive(x_11); +if (x_25 == 0) { return x_11; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_11, 0); -x_26 = lean_ctor_get(x_11, 1); +lean_object* x_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_11, 0); +x_27 = lean_ctor_get(x_11, 1); +lean_inc(x_27); lean_inc(x_26); -lean_inc(x_25); lean_dec(x_11); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +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; } } } diff --git a/stage0/stdlib/Lean/Elab/Tactic/SolveByElim.c b/stage0/stdlib/Lean/Elab/Tactic/SolveByElim.c index 5419468e2709..d1ef9abb65ef 100644 --- a/stage0/stdlib/Lean/Elab/Tactic/SolveByElim.c +++ b/stage0/stdlib/Lean/Elab/Tactic/SolveByElim.c @@ -70,7 +70,7 @@ static lean_object* l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__9; static lean_object* l_Lean_Elab_Tactic_SolveByElim_parseUsing___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Tactic_SolveByElim_elabConfig___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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SolveByElim_parseArgs(lean_object*); -lean_object* l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_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_EXPORT lean_object* l_panic___at_Lean_Elab_Tactic_SolveByElim_parseArgs___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_SolveByElim_evalApplyAssumption___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Sum_getRight_x3f___rarg(lean_object*); @@ -2588,38 +2588,39 @@ lean_inc(x_4); x_12 = l_Lean_Elab_Term_elabTermEnsuringType(x_1, x_2, x_11, x_11, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; +lean_object* x_13; lean_object* x_14; uint8_t x_15; uint8_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); lean_dec(x_12); -x_15 = 0; +x_15 = 1; +x_16 = 0; lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); -x_16 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_15, x_4, x_5, x_6, x_7, x_8, x_9, x_14); -if (lean_obj_tag(x_16) == 0) +x_17 = l_Lean_Elab_Term_synthesizeSyntheticMVars(x_15, x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_14); +if (lean_obj_tag(x_17) == 0) { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_16, 1); -lean_inc(x_17); -lean_dec(x_16); -x_18 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_17); +lean_object* x_18; lean_object* x_19; +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +lean_dec(x_17); +x_19 = l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(x_13, x_4, x_5, x_6, x_7, x_8, x_9, x_18); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -return x_18; +return x_19; } else { -uint8_t x_19; +uint8_t x_20; lean_dec(x_13); lean_dec(x_9); lean_dec(x_8); @@ -2627,52 +2628,52 @@ lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_19 = !lean_is_exclusive(x_16); -if (x_19 == 0) +x_20 = !lean_is_exclusive(x_17); +if (x_20 == 0) { -return x_16; +return x_17; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_16, 0); -x_21 = lean_ctor_get(x_16, 1); +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_17, 0); +x_22 = lean_ctor_get(x_17, 1); +lean_inc(x_22); lean_inc(x_21); -lean_inc(x_20); -lean_dec(x_16); -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_17); +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_9); lean_dec(x_8); lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); -x_23 = !lean_is_exclusive(x_12); -if (x_23 == 0) +x_24 = !lean_is_exclusive(x_12); +if (x_24 == 0) { return x_12; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_12, 0); -x_25 = lean_ctor_get(x_12, 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_12); -x_26 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_26, 0, x_24); -lean_ctor_set(x_26, 1, x_25); -return x_26; +x_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; } } } @@ -2928,7 +2929,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -2940,85 +2941,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_SolveByElim_elabConfig___la lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__7; -x_24 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__7; +x_22 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_SolveByElim_evalUnsafe____x40_Lean_Elab_Tactic_SolveByElim___hyg_8_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_SolveByElim_evalUnsafe____x40_Lean_Elab_Tactic_SolveByElim___hyg_8_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__17; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__17; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } } } @@ -3268,7 +3266,7 @@ uint8_t x_9; x_9 = l_Lean_Syntax_isNone(x_1); if (x_9 == 0) { -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t 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; lean_object* x_24; lean_object* x_25; lean_object* x_26; +lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; lean_object* x_18; 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_10 = lean_unsigned_to_nat(0u); x_11 = l_Lean_Syntax_getArg(x_1, x_10); x_12 = lean_unsigned_to_nat(3u); @@ -3280,85 +3278,82 @@ x_16 = lean_alloc_closure((void*)(l_Lean_Elab_Tactic_SolveByElim_elabConfig___la lean_closure_set(x_16, 0, x_13); lean_closure_set(x_16, 1, x_15); lean_closure_set(x_16, 2, x_14); -x_17 = 0; -x_18 = 1; -x_19 = lean_box(x_17); -x_20 = lean_box(x_18); -x_21 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 10, 3); -lean_closure_set(x_21, 0, x_16); -lean_closure_set(x_21, 1, x_19); -lean_closure_set(x_21, 2, x_20); -x_22 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__1), 8, 1); -lean_closure_set(x_22, 0, x_21); -x_23 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__7; -x_24 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__8; -x_25 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); -lean_closure_set(x_25, 0, x_23); -lean_closure_set(x_25, 1, x_24); -lean_closure_set(x_25, 2, x_22); +x_17 = 1; +x_18 = lean_box(x_17); +x_19 = lean_alloc_closure((void*)(l___private_Lean_Elab_SyntheticMVars_0__Lean_Elab_Term_withSynthesizeImp___rarg___boxed), 9, 2); +lean_closure_set(x_19, 0, x_16); +lean_closure_set(x_19, 1, x_18); +x_20 = lean_alloc_closure((void*)(l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__1), 8, 1); +lean_closure_set(x_20, 0, x_19); +x_21 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__7; +x_22 = l_Lean_Elab_Tactic_SolveByElim_elabConfig___closed__8; +x_23 = lean_alloc_closure((void*)(l_Lean_Meta_withLCtx___at___private_Lean_Elab_Binders_0__Lean_Elab_Term_FunBinders_elabFunBinderViews___spec__3___rarg), 10, 3); +lean_closure_set(x_23, 0, x_21); +lean_closure_set(x_23, 1, x_22); +lean_closure_set(x_23, 2, x_20); lean_inc(x_7); lean_inc(x_6); lean_inc(x_5); lean_inc(x_4); lean_inc(x_3); lean_inc(x_2); -x_26 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__13(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -if (lean_obj_tag(x_26) == 0) +x_24 = l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___at_Lean_Elab_Tactic_SolveByElim_elabConfig___spec__13(x_23, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +if (lean_obj_tag(x_24) == 0) { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -x_27 = lean_ctor_get(x_26, 0); -lean_inc(x_27); -x_28 = lean_ctor_get(x_26, 1); -lean_inc(x_28); -lean_dec(x_26); -x_29 = l_Lean_Elab_Tactic_SolveByElim_evalUnsafe____x40_Lean_Elab_Tactic_SolveByElim___hyg_193_(x_27, x_2, x_3, x_4, x_5, x_6, x_7, x_28); +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_SolveByElim_evalUnsafe____x40_Lean_Elab_Tactic_SolveByElim___hyg_193_(x_25, x_2, x_3, x_4, x_5, x_6, x_7, x_26); lean_dec(x_3); lean_dec(x_2); -return x_29; +return x_27; } else { -uint8_t x_30; +uint8_t x_28; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_30 = !lean_is_exclusive(x_26); -if (x_30 == 0) +x_28 = !lean_is_exclusive(x_24); +if (x_28 == 0) { -return x_26; +return x_24; } else { -lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_31 = lean_ctor_get(x_26, 0); -x_32 = lean_ctor_get(x_26, 1); -lean_inc(x_32); -lean_inc(x_31); -lean_dec(x_26); -x_33 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -return x_33; +lean_object* x_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 { -lean_object* x_34; lean_object* x_35; +lean_object* x_32; lean_object* x_33; lean_dec(x_7); lean_dec(x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_34 = l_Lean_Elab_Tactic_SolveByElim_elabApplyRulesConfig___closed__4; -x_35 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_35, 0, x_34); -lean_ctor_set(x_35, 1, x_8); -return x_35; +x_32 = l_Lean_Elab_Tactic_SolveByElim_elabApplyRulesConfig___closed__4; +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_8); +return x_33; } } } diff --git a/stage0/stdlib/Lean/Elab/Term.c b/stage0/stdlib/Lean/Elab/Term.c index 2576d3c0f9c7..bbb141b764eb 100644 --- a/stage0/stdlib/Lean/Elab/Term.c +++ b/stage0/stdlib/Lean/Elab/Term.c @@ -73,7 +73,6 @@ LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__13(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getMVarErrorInfo_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_collectUnassignedMVars___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___closed__2; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__26___boxed(lean_object*, lean_object*, lean_object*); @@ -112,11 +111,11 @@ lean_object* lean_private_to_user_name(lean_object*); lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_shouldVisit(lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__5; LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Elab_Term_isLetRecAuxMVar___spec__3(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Level_elabLevel(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM; LEAN_EXPORT lean_object* l___private_Lean_Elab_InfoTree_Main_0__Lean_Elab_withSavedPartialInfoContext___at_Lean_Elab_Term_withDeclName___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 uint8_t l_Lean_Elab_Term_isLetRecAuxMVar___lambda__1(lean_object*, lean_object*); -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_expandDeclId___spec__12(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -202,7 +201,6 @@ LEAN_EXPORT lean_object* l_Lean_getDelayedMVarAssignment_x3f___at_Lean_Elab_Term lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_elabTermEnsuringType(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -228,7 +226,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerSyntheticMVar(lean_object*, le LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1___boxed(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_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___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_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___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_throwError___at_Lean_Elab_Term_mkConst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isNoImplicitLambda___closed__2; lean_object* l_String_removeLeadingSpaces(lean_object*); 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*); @@ -287,7 +284,6 @@ static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__30(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withoutModifyingStateWithInfoAndMessages___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instHashableCacheKey; -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars___closed__2; lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); @@ -347,7 +343,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_saveContext___boxed(lean_object*, lean 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); -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_dropTermParens___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMonadBacktrackSavedStateTermElabM___closed__3; @@ -368,7 +363,6 @@ static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expa LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__5(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVarError___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError___closed__3; -LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__16; lean_object* lean_io_get_num_heartbeats(lean_object*); uint8_t l_Lean_Syntax_isOfKind(lean_object*, lean_object*); @@ -395,7 +389,6 @@ static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__25; 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_Cache_post___default; -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*); lean_object* l_Lean_PersistentArray_toList___rarg(lean_object*); @@ -413,6 +406,7 @@ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaE LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMonadMacroAdapterTermElabM___lambda__2(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_withDeclName___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* l_Lean_Elab_expandMacroImpl_x3f___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9; uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_instInhabitedTermElabM(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkCoe___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -428,6 +422,7 @@ static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__15; LEAN_EXPORT uint8_t l_Lean_Elab_Term_levelMVarToParam___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentArray_mapMAux___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visit___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__18___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___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_PersistentArray_forMAux___at_Lean_Elab_Term_instMetaEvalTermElabM___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_SavedState_restoreFull(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__51___closed__3; @@ -486,9 +481,11 @@ uint8_t l_Lean_Expr_hasMVar(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedCacheKey; uint8_t l_Lean_isReservedName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withMacroExpansionInfo___at_Lean_Elab_Term_withMacroExpansion___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___closed__4; LEAN_EXPORT lean_object* l_Lean_Elab_Term_TermElabM_run(lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Elab_Term_0__Lean_Elab_Term_isTypeAscription(lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instantiateMVars___at_Lean_Elab_Term_MVarErrorInfo_logError___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_logAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -525,10 +522,10 @@ LEAN_EXPORT lean_object* l_Lean_profileitM___at___private_Lean_Elab_Term_0__Lean LEAN_EXPORT lean_object* l_Lean_Elab_Term_withSavedContext___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_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_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*); +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_2009____closed__1; static lean_object* l_Lean_Elab_Term_instMonadTermElabM___closed__4; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__3___closed__3; @@ -553,7 +550,6 @@ static lean_object* l_Lean_Elab_Term_withoutHeedElabAsElim___rarg___closed__1; static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___closed__2; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withExpectedType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_local_ctx_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_registerMVarErrorImplicitArgInfo(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -586,7 +582,6 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea lean_object* l_Lean_mkPrivateName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at_Lean_Elab_Term_addDotCompletionInfo___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__6; -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Elab_Term_0__Lean_Elab_Term_withoutModifyingStateWithInfoAndMessagesImpl___spec__11(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Elab_Term_mkCoe___spec__1___lambda__4(lean_object*, lean_object*, 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 lean_object* l_Lean_Elab_throwPostpone___at_Lean_Elab_Term_tryPostpone___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -608,7 +603,6 @@ LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_v static lean_object* l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__2; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__9(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_Term_getLevelNames(lean_object*); -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__1; @@ -616,7 +610,6 @@ LEAN_EXPORT uint8_t l_Lean_PersistentArray_anyMAux___at_Lean_Elab_Term_addAutoBo LEAN_EXPORT lean_object* l_Lean_Elab_Tactic_instInhabitedCache; static lean_object* l_Lean_Elab_Term_mkTypeMismatchError___closed__6; lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortCommand___at_Lean_Elab_Term_ensureNoUnassignedMVars___spec__1___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withoutErrToSorryImp___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Elab_Term_traceAtCmdPos___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -644,7 +637,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_ppGoal(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfHasMVars_x3f___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_forInUnsafe_loop___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__10___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__3___closed__3; -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_withSaveParentDeclInfoContext___at_Lean_Elab_Term_withDeclName___spec__1___rarg___lambda__1(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__5; static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__5; @@ -670,6 +662,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_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; 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 +673,6 @@ static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13936___ static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabImplicitLambdaAux___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_addAutoBoundImplicits___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_resolveId_x3f___lambda__2___closed__4; -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_decorateErrorMessageWithLambdaImplicitVars(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withAuxDecl(lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Elab_Term_logUnassignedUsingErrorInfos___spec__6(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -767,6 +759,7 @@ 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_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3; 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); @@ -844,7 +837,6 @@ lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_addTraceNode___at_Lean_Elab_Term_mkCoe___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_extractMacroScopes(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwPostpone___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_mkFreshBinderName___rarg___closed__2; lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___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*); @@ -881,6 +873,7 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expan LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___lambda__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyResult___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1; static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__30; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Elab_Term_resolveLocalName___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedLetRecToLift___closed__1; @@ -908,6 +901,7 @@ LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Elab_Term_MVarErrorInfo_logError uint8_t l_Lean_Syntax_matchesNull(lean_object*, lean_object*); static lean_object* l_Lean_Elab_checkNotAlreadyDeclared___at_Lean_Elab_Term_expandDeclId___spec__6___lambda__1___closed__3; lean_object* l_Lean_Core_setMessageLog(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1; static lean_object* l_Lean_Elab_Tactic_instHashableCacheKey___closed__1; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__25(lean_object*, lean_object*, size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Elab_withSaveInfoContext___at_Lean_Elab_Term_withoutModifyingElabMetaStateWithInfo___spec__1___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1004,6 +998,7 @@ 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*); +LEAN_EXPORT lean_object* l_List_foldlM___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*, lean_object*, lean_object*); 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; @@ -1041,6 +1036,7 @@ static lean_object* l_Lean_Elab_Term_resolveName___closed__4; LEAN_EXPORT 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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkAuxName___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7; LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_Term_addTermInfo___spec__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___rarg___closed__1; LEAN_EXPORT lean_object* l_Array_findSomeRevM_x3f_find___at_Lean_Elab_Term_resolveLocalName___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1071,11 +1067,13 @@ 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_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2; 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; +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_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; +LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_isExplicit___closed__2; static lean_object* l_Lean_Elab_Term_withExpectedType___closed__1; static lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___rarg___closed__1; @@ -1214,7 +1212,6 @@ uint8_t l_Lean_Syntax_isNone(lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_applyAttributes(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitApp___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__20(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_postponeElabTermCore___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_annotation_x3f(lean_object*, lean_object*); @@ -1322,6 +1319,7 @@ LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImpli static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__9; lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__15(lean_object*, lean_object*, size_t, size_t); +LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_throwMVarError___spec__2___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_Term_observing___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -1333,7 +1331,6 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2(lea LEAN_EXPORT lean_object* l_Lean_Elab_withSaveParentDeclInfoContext___at_Lean_Elab_Term_withDeclName___spec__1(lean_object*); uint8_t l_Lean_Name_quickCmp(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwMaxRecDepthAt___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__14(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_instMetaEvalTermElabM___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_instInhabitedTermElabM___closed__2; @@ -1399,6 +1396,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkSyntheti 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*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_tryPostponeIfNoneOrMVar(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_List_foldlM___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_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__32___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13936____closed__12; 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*); @@ -1446,6 +1444,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_logException___at___private_Lean_Elab_Term_ LEAN_EXPORT lean_object* l_Lean_Elab_Term_instToStringLVal(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwAbortTerm___at_Lean_Elab_Term_ensureType___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFnsAux___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_Term_isLetRecAuxMVar___lambda__3___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_Term_withDeclName___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*); uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); @@ -1476,7 +1475,6 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Elab_Term_logUnassi lean_object* l_Lean_realizeGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT 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*); 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; LEAN_EXPORT lean_object* l_Lean_Elab_Term_instMetaEvalTermElabM___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*); @@ -1494,6 +1492,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkFreshIdent___rarg(lean_object*, lean LEAN_EXPORT lean_object* l_Lean_Elab_throwUnsupportedSyntax___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expandDeclId___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___spec__1___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4; LEAN_EXPORT lean_object* l___private_Lean_MetavarContext_0__Lean_DependsOn_dep_visitMain___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__19(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoLeaf___at___private_Lean_Elab_Term_0__Lean_Elab_Term_applyAttributesCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_pushInfoTree___at_Lean_Elab_Term_addDotCompletionInfo___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1507,6 +1506,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_Term_SavedState_restore___boxed(lean_object static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabTermAux___lambda__7___closed__2; static lean_object* l_Lean_Elab_Term_expandDeclId___closed__2; static lean_object* l_Lean_Elab_Term_initFn____x40_Lean_Elab_Term___hyg_13936____closed__2; +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_elabUsingElabFns___closed__1; static lean_object* l_Lean_Elab_Term_synthesizeInstMVarCore___lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_Term_SavedState_restore(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1560,6 +1560,7 @@ LEAN_EXPORT lean_object* l_Lean_Elab_throwAlreadyDeclaredUniverseLevel___at_Lean static lean_object* l_Array_forInUnsafe_loop___at_Lean_Elab_Term_expandDeclId___spec__4___closed__5; static lean_object* l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__8; static lean_object* l_Lean_Elab_Term_mkExplicitBinder___closed__8; +static lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8; LEAN_EXPORT uint8_t l_Lean_Elab_Term_Context_ignoreTCFailures___default; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Elab_Term_addAutoBoundImplicits_go___spec__47___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Elab_Term_collectUnassignedMVars_go___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -1590,7 +1591,6 @@ static lean_object* l_Lean_Elab_Term_ContainsPendingMVar_visit___closed__1; uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_expandDeclId___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Message_toString(lean_object*, uint8_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_Term_withDeclName___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* _init_l_Lean_Elab_Term_instInhabitedSyntheticMVarKind() { _start: @@ -53166,325 +53166,7 @@ lean_dec(x_2); return x_10; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; -x_9 = lean_ctor_get(x_6, 5); -x_10 = lean_ctor_get(x_2, 2); -lean_inc(x_10); -lean_inc(x_10); -x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); -x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); -x_13 = lean_ctor_get(x_12, 0); -lean_inc(x_13); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); -lean_dec(x_2); -x_16 = !lean_is_exclusive(x_15); -if (x_16 == 0) -{ -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_15, 0); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_17); -lean_ctor_set_tag(x_15, 1); -lean_ctor_set(x_15, 0, x_18); -return x_15; -} -else -{ -lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_19 = lean_ctor_get(x_15, 0); -x_20 = lean_ctor_get(x_15, 1); -lean_inc(x_20); -lean_inc(x_19); -lean_dec(x_15); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_11); -lean_ctor_set(x_21, 1, x_19); -x_22 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_20); -return x_22; -} -} -} -static lean_object* _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("unknown constant '", 18); -return x_1; -} -} -static lean_object* _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; uint8_t x_10; -x_9 = lean_st_ref_get(x_7, x_8); -x_10 = !lean_is_exclusive(x_9); -if (x_10 == 0) -{ -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_9, 0); -x_12 = lean_ctor_get(x_9, 1); -x_13 = lean_ctor_get(x_11, 0); -lean_inc(x_13); -lean_dec(x_11); -lean_inc(x_1); -x_14 = lean_environment_find(x_13, x_1); -if (lean_obj_tag(x_14) == 0) -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; -lean_free_object(x_9); -x_15 = lean_box(0); -x_16 = l_Lean_Expr_const___override(x_1, x_15); -x_17 = l_Lean_MessageData_ofExpr(x_16); -x_18 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2; -x_19 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_19, 0, x_18); -lean_ctor_set(x_19, 1, x_17); -x_20 = l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___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_Lean_Elab_Term_mkConst___spec__2(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_12); -return x_22; -} -else -{ -lean_object* x_23; -lean_dec(x_2); -lean_dec(x_1); -x_23 = lean_ctor_get(x_14, 0); -lean_inc(x_23); -lean_dec(x_14); -lean_ctor_set(x_9, 0, x_23); -return x_9; -} -} -else -{ -lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; -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_ctor_get(x_24, 0); -lean_inc(x_26); -lean_dec(x_24); -lean_inc(x_1); -x_27 = lean_environment_find(x_26, x_1); -if (lean_obj_tag(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; -x_28 = lean_box(0); -x_29 = l_Lean_Expr_const___override(x_1, x_28); -x_30 = l_Lean_MessageData_ofExpr(x_29); -x_31 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2; -x_32 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_32, 0, x_31); -lean_ctor_set(x_32, 1, x_30); -x_33 = l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__4; -x_34 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_34, 0, x_32); -lean_ctor_set(x_34, 1, x_33); -x_35 = l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_25); -return x_35; -} -else -{ -lean_object* x_36; lean_object* x_37; -lean_dec(x_2); -lean_dec(x_1); -x_36 = lean_ctor_get(x_27, 0); -lean_inc(x_36); -lean_dec(x_27); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, x_25); -return x_37; -} -} -} -} -static lean_object* _init_l_Lean_Elab_Term_mkConst___closed__1() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("too many explicit universe levels for '", 39); -return x_1; -} -} -static lean_object* _init_l_Lean_Elab_Term_mkConst___closed__2() { -_start: -{ -lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Elab_Term_mkConst___closed__1; -x_2 = l_Lean_stringToMessageData(x_1); -return x_2; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_3); -lean_inc(x_1); -x_10 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); -if (lean_obj_tag(x_10) == 0) -{ -lean_object* x_11; 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_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = l_Lean_ConstantInfo_levelParams(x_11); -lean_dec(x_11); -x_14 = lean_unsigned_to_nat(0u); -x_15 = l_List_lengthTRAux___rarg(x_13, x_14); -lean_dec(x_13); -x_16 = l_List_lengthTRAux___rarg(x_2, x_14); -x_17 = lean_nat_dec_lt(x_15, x_16); -if (x_17 == 0) -{ -lean_object* x_18; lean_object* x_19; uint8_t x_20; -lean_dec(x_3); -x_18 = lean_nat_sub(x_15, x_16); -lean_dec(x_16); -lean_dec(x_15); -x_19 = l_Lean_Meta_mkFreshLevelMVars(x_18, x_5, x_6, x_7, x_8, x_12); -x_20 = !lean_is_exclusive(x_19); -if (x_20 == 0) -{ -lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_21 = lean_ctor_get(x_19, 0); -x_22 = l_List_appendTR___rarg(x_2, x_21); -x_23 = l_Lean_Expr_const___override(x_1, x_22); -lean_ctor_set(x_19, 0, x_23); -return x_19; -} -else -{ -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_19, 0); -x_25 = lean_ctor_get(x_19, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_19); -x_26 = l_List_appendTR___rarg(x_2, x_24); -x_27 = l_Lean_Expr_const___override(x_1, x_26); -x_28 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_28, 0, x_27); -lean_ctor_set(x_28, 1, x_25); -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_dec(x_16); -lean_dec(x_15); -lean_dec(x_2); -x_29 = l_Lean_MessageData_ofName(x_1); -x_30 = l_Lean_Elab_Term_mkConst___closed__2; -x_31 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_29); -x_32 = l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__4; -x_33 = lean_alloc_ctor(7, 2, 0); -lean_ctor_set(x_33, 0, x_31); -lean_ctor_set(x_33, 1, x_32); -x_34 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_33, x_3, x_4, x_5, x_6, x_7, x_8, x_12); -lean_dec(x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -return x_34; -} -} -else -{ -uint8_t x_35; -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_is_exclusive(x_10); -if (x_35 == 0) -{ -return x_10; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_36 = lean_ctor_get(x_10, 0); -x_37 = lean_ctor_get(x_10, 1); -lean_inc(x_37); -lean_inc(x_36); -lean_dec(x_10); -x_38 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_38, 0, x_36); -lean_ctor_set(x_38, 1, x_37); -return x_38; -} -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { -_start: -{ -lean_object* x_10; -x_10 = l_Lean_Elab_Term_mkConst(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_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -53507,7 +53189,7 @@ return x_15; } } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1() { _start: { lean_object* x_1; @@ -53515,7 +53197,7 @@ x_1 = l_Lean_Linter_linter_deprecated; return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2() { _start: { lean_object* x_1; @@ -53523,7 +53205,7 @@ x_1 = l_Lean_Linter_deprecatedAttr; return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3() { _start: { lean_object* x_1; @@ -53531,7 +53213,7 @@ x_1 = lean_mk_string_from_bytes("Linter", 6); return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4() { _start: { lean_object* x_1; @@ -53539,18 +53221,18 @@ x_1 = lean_mk_string_from_bytes("deprecatedAttr", 14); return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_Term_mkTermElabAttributeUnsafe___closed__5; -x_2 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3; -x_3 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4; +x_2 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3; +x_3 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6() { _start: { lean_object* x_1; @@ -53558,7 +53240,7 @@ x_1 = lean_mk_string_from_bytes("`", 1); return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7() { _start: { lean_object* x_1; @@ -53566,7 +53248,7 @@ x_1 = lean_mk_string_from_bytes("` has been deprecated", 21); return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8() { _start: { lean_object* x_1; @@ -53574,7 +53256,7 @@ x_1 = lean_mk_string_from_bytes("` has been deprecated, use `", 28); return x_1; } } -static lean_object* _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9() { +static lean_object* _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9() { _start: { lean_object* x_1; @@ -53582,12 +53264,12 @@ x_1 = lean_mk_string_from_bytes("` instead", 9); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { lean_object* x_9; lean_object* x_10; uint8_t x_11; x_9 = lean_ctor_get(x_6, 2); -x_10 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1; +x_10 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1; x_11 = l_Lean_Linter_getLinterValue(x_10, x_9); if (x_11 == 0) { @@ -53613,7 +53295,7 @@ x_18 = lean_ctor_get(x_16, 0); lean_inc(x_18); lean_dec(x_16); x_19 = l_Lean_Linter_instInhabitedDeprecationEntry; -x_20 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2; +x_20 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2; lean_inc(x_1); x_21 = l_Lean_ParametricAttribute_getParam_x3f___rarg(x_19, x_20, x_18, x_1); if (lean_obj_tag(x_21) == 0) @@ -53644,20 +53326,20 @@ if (lean_obj_tag(x_25) == 0) uint8_t x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; x_26 = 1; x_27 = l_Lean_Name_toString(x_1, x_26); -x_28 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; +x_28 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6; x_29 = lean_string_append(x_28, x_27); lean_dec(x_27); -x_30 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7; +x_30 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7; x_31 = lean_string_append(x_29, x_30); x_32 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_32, 0, x_31); x_33 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_33, 0, x_32); -x_34 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_34 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_35 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_35, 0, x_34); lean_ctor_set(x_35, 1, x_33); -x_36 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +x_36 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_35, x_2, x_3, x_4, x_5, x_6, x_7, x_17); return x_36; } else @@ -53668,25 +53350,25 @@ lean_inc(x_37); lean_dec(x_25); x_38 = 1; x_39 = l_Lean_Name_toString(x_1, x_38); -x_40 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; +x_40 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6; x_41 = lean_string_append(x_40, x_39); lean_dec(x_39); -x_42 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8; +x_42 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8; x_43 = lean_string_append(x_41, x_42); x_44 = l_Lean_Name_toString(x_37, x_38); x_45 = lean_string_append(x_43, x_44); lean_dec(x_44); -x_46 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9; +x_46 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9; x_47 = lean_string_append(x_45, x_46); x_48 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_48, 0, x_47); x_49 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_49, 0, x_48); -x_50 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_50 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_51 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_51, 0, x_50); lean_ctor_set(x_51, 1, x_49); -x_52 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_51, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +x_52 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_51, x_2, x_3, x_4, x_5, x_6, x_7, x_17); return x_52; } } @@ -53702,11 +53384,11 @@ x_54 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_54, 0, x_53); x_55 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_55, 0, x_54); -x_56 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_56 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_57 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_57, 0, x_56); lean_ctor_set(x_57, 1, x_55); -x_58 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_17); +x_58 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_57, x_2, x_3, x_4, x_5, x_6, x_7, x_17); return x_58; } } @@ -53723,7 +53405,7 @@ x_61 = lean_ctor_get(x_59, 0); lean_inc(x_61); lean_dec(x_59); x_62 = l_Lean_Linter_instInhabitedDeprecationEntry; -x_63 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2; +x_63 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2; lean_inc(x_1); x_64 = l_Lean_ParametricAttribute_getParam_x3f___rarg(x_62, x_63, x_61, x_1); if (lean_obj_tag(x_64) == 0) @@ -53755,20 +53437,20 @@ if (lean_obj_tag(x_69) == 0) 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; lean_object* x_79; lean_object* x_80; x_70 = 1; x_71 = l_Lean_Name_toString(x_1, x_70); -x_72 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; +x_72 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6; x_73 = lean_string_append(x_72, x_71); lean_dec(x_71); -x_74 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7; +x_74 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7; x_75 = lean_string_append(x_73, x_74); x_76 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_76, 0, x_75); x_77 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_77, 0, x_76); -x_78 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_78 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_79 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_79, 0, x_78); lean_ctor_set(x_79, 1, x_77); -x_80 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_79, x_2, x_3, x_4, x_5, x_6, x_7, x_60); +x_80 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_79, x_2, x_3, x_4, x_5, x_6, x_7, x_60); return x_80; } else @@ -53779,25 +53461,25 @@ lean_inc(x_81); lean_dec(x_69); x_82 = 1; x_83 = l_Lean_Name_toString(x_1, x_82); -x_84 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6; +x_84 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6; x_85 = lean_string_append(x_84, x_83); lean_dec(x_83); -x_86 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8; +x_86 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8; x_87 = lean_string_append(x_85, x_86); x_88 = l_Lean_Name_toString(x_81, x_82); x_89 = lean_string_append(x_87, x_88); lean_dec(x_88); -x_90 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9; +x_90 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9; x_91 = lean_string_append(x_89, x_90); x_92 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_92, 0, x_91); x_93 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_93, 0, x_92); -x_94 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_94 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_95 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_95, 0, x_94); lean_ctor_set(x_95, 1, x_93); -x_96 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_60); +x_96 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_95, x_2, x_3, x_4, x_5, x_6, x_7, x_60); return x_96; } } @@ -53813,11 +53495,11 @@ x_98 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_98, 0, x_97); x_99 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_99, 0, x_98); -x_100 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5; +x_100 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5; x_101 = lean_alloc_ctor(8, 2, 0); lean_ctor_set(x_101, 0, x_100); lean_ctor_set(x_101, 1, x_99); -x_102 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_60); +x_102 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_101, x_2, x_3, x_4, x_5, x_6, x_7, x_60); return x_102; } } @@ -53825,7 +53507,358 @@ return x_102; } } } -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; +x_9 = lean_ctor_get(x_6, 5); +x_10 = lean_ctor_get(x_2, 2); +lean_inc(x_10); +lean_inc(x_10); +x_11 = l_Lean_Elab_getBetterRef(x_9, x_10); +x_12 = l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(x_1, x_4, x_5, x_6, x_7, x_8); +x_13 = lean_ctor_get(x_12, 0); +lean_inc(x_13); +x_14 = lean_ctor_get(x_12, 1); +lean_inc(x_14); +lean_dec(x_12); +x_15 = l_Lean_Elab_addMacroStack___at_Lean_Elab_Term_instAddErrorMessageContextTermElabM___spec__1(x_13, x_10, x_2, x_3, x_4, x_5, x_6, x_7, x_14); +lean_dec(x_2); +x_16 = !lean_is_exclusive(x_15); +if (x_16 == 0) +{ +lean_object* x_17; lean_object* x_18; +x_17 = lean_ctor_get(x_15, 0); +x_18 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_18, 0, x_11); +lean_ctor_set(x_18, 1, x_17); +lean_ctor_set_tag(x_15, 1); +lean_ctor_set(x_15, 0, x_18); +return x_15; +} +else +{ +lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +x_19 = lean_ctor_get(x_15, 0); +x_20 = lean_ctor_get(x_15, 1); +lean_inc(x_20); +lean_inc(x_19); +lean_dec(x_15); +x_21 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_21, 0, x_11); +lean_ctor_set(x_21, 1, x_19); +x_22 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_22, 0, x_21); +lean_ctor_set(x_22, 1, x_20); +return x_22; +} +} +} +static lean_object* _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("unknown constant '", 18); +return x_1; +} +} +static lean_object* _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___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; uint8_t x_10; +x_9 = lean_st_ref_get(x_7, x_8); +x_10 = !lean_is_exclusive(x_9); +if (x_10 == 0) +{ +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; +x_11 = lean_ctor_get(x_9, 0); +x_12 = lean_ctor_get(x_9, 1); +x_13 = lean_ctor_get(x_11, 0); +lean_inc(x_13); +lean_dec(x_11); +lean_inc(x_1); +x_14 = lean_environment_find(x_13, x_1); +if (lean_obj_tag(x_14) == 0) +{ +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; +lean_free_object(x_9); +x_15 = lean_box(0); +x_16 = l_Lean_Expr_const___override(x_1, x_15); +x_17 = l_Lean_MessageData_ofExpr(x_16); +x_18 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2; +x_19 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_19, 0, x_18); +lean_ctor_set(x_19, 1, x_17); +x_20 = l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___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_Lean_Elab_Term_mkConst___spec__4(x_21, x_2, x_3, x_4, x_5, x_6, x_7, x_12); +return x_22; +} +else +{ +lean_object* x_23; +lean_dec(x_2); +lean_dec(x_1); +x_23 = lean_ctor_get(x_14, 0); +lean_inc(x_23); +lean_dec(x_14); +lean_ctor_set(x_9, 0, x_23); +return x_9; +} +} +else +{ +lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; +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_ctor_get(x_24, 0); +lean_inc(x_26); +lean_dec(x_24); +lean_inc(x_1); +x_27 = lean_environment_find(x_26, x_1); +if (lean_obj_tag(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; +x_28 = lean_box(0); +x_29 = l_Lean_Expr_const___override(x_1, x_28); +x_30 = l_Lean_MessageData_ofExpr(x_29); +x_31 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2; +x_32 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_30); +x_33 = l_Lean_Elab_Term_MVarErrorInfo_logError_addArgName___closed__4; +x_34 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_34, 0, x_32); +lean_ctor_set(x_34, 1, x_33); +x_35 = l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4(x_34, x_2, x_3, x_4, x_5, x_6, x_7, x_25); +return x_35; +} +else +{ +lean_object* x_36; lean_object* x_37; +lean_dec(x_2); +lean_dec(x_1); +x_36 = lean_ctor_get(x_27, 0); +lean_inc(x_36); +lean_dec(x_27); +x_37 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_37, 0, x_36); +lean_ctor_set(x_37, 1, x_25); +return x_37; +} +} +} +} +static lean_object* _init_l_Lean_Elab_Term_mkConst___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("too many explicit universe levels for '", 39); +return x_1; +} +} +static lean_object* _init_l_Lean_Elab_Term_mkConst___closed__2() { +_start: +{ +lean_object* x_1; lean_object* x_2; +x_1 = l_Lean_Elab_Term_mkConst___closed__1; +x_2 = l_Lean_stringToMessageData(x_1); +return x_2; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst(lean_object* x_1, lean_object* x_2, lean_object* x_3, 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_inc(x_1); +x_10 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = lean_ctor_get(x_10, 1); +lean_inc(x_11); +lean_dec(x_10); +lean_inc(x_3); +lean_inc(x_1); +x_12 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_1, x_3, x_4, x_5, x_6, x_7, x_8, 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; uint8_t x_19; +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_ConstantInfo_levelParams(x_13); +lean_dec(x_13); +x_16 = lean_unsigned_to_nat(0u); +x_17 = l_List_lengthTRAux___rarg(x_15, x_16); +lean_dec(x_15); +x_18 = l_List_lengthTRAux___rarg(x_2, x_16); +x_19 = lean_nat_dec_lt(x_17, x_18); +if (x_19 == 0) +{ +lean_object* x_20; lean_object* x_21; uint8_t x_22; +lean_dec(x_3); +x_20 = lean_nat_sub(x_17, x_18); +lean_dec(x_18); +lean_dec(x_17); +x_21 = l_Lean_Meta_mkFreshLevelMVars(x_20, x_5, x_6, x_7, x_8, x_14); +x_22 = !lean_is_exclusive(x_21); +if (x_22 == 0) +{ +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_21, 0); +x_24 = l_List_appendTR___rarg(x_2, x_23); +x_25 = l_Lean_Expr_const___override(x_1, x_24); +lean_ctor_set(x_21, 0, x_25); +return x_21; +} +else +{ +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; +x_26 = lean_ctor_get(x_21, 0); +x_27 = lean_ctor_get(x_21, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_21); +x_28 = l_List_appendTR___rarg(x_2, x_26); +x_29 = l_Lean_Expr_const___override(x_1, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 0, x_29); +lean_ctor_set(x_30, 1, x_27); +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_dec(x_18); +lean_dec(x_17); +lean_dec(x_2); +x_31 = l_Lean_MessageData_ofName(x_1); +x_32 = l_Lean_Elab_Term_mkConst___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_Term_MVarErrorInfo_logError_addArgName___closed__4; +x_35 = lean_alloc_ctor(7, 2, 0); +lean_ctor_set(x_35, 0, x_33); +lean_ctor_set(x_35, 1, x_34); +x_36 = l_Lean_throwError___at_Lean_Elab_Term_tryPostponeIfHasMVars___spec__1(x_35, x_3, x_4, x_5, x_6, x_7, x_8, x_14); +lean_dec(x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +return x_36; +} +} +else +{ +uint8_t x_37; +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_37 = !lean_is_exclusive(x_12); +if (x_37 == 0) +{ +return x_12; +} +else +{ +lean_object* x_38; lean_object* x_39; lean_object* x_40; +x_38 = lean_ctor_get(x_12, 0); +x_39 = lean_ctor_get(x_12, 1); +lean_inc(x_39); +lean_inc(x_38); +lean_dec(x_12); +x_40 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_40, 0, x_38); +lean_ctor_set(x_40, 1, x_39); +return x_40; +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_logWarning___at_Lean_Elab_Term_mkConst___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_2); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_throwError___at_Lean_Elab_Term_mkConst___spec__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { +_start: +{ +lean_object* x_9; +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +lean_dec(x_7); +lean_dec(x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_3); +return x_9; +} +} +LEAN_EXPORT lean_object* l_Lean_Elab_Term_mkConst___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9) { +_start: +{ +lean_object* x_10; +x_10 = l_Lean_Elab_Term_mkConst(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_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { if (lean_obj_tag(x_3) == 0) @@ -53853,40 +53886,35 @@ x_13 = lean_ctor_get(x_3, 0); x_14 = !lean_is_exclusive(x_13); if (x_14 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; +lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; x_15 = lean_ctor_get(x_3, 1); x_16 = lean_ctor_get(x_13, 0); x_17 = lean_ctor_get(x_13, 1); -lean_inc(x_16); -x_18 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_16, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_19 = lean_ctor_get(x_18, 1); -lean_inc(x_19); -lean_dec(x_18); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_4); lean_inc(x_1); -x_20 = l_Lean_Elab_Term_mkConst(x_16, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_19); -if (lean_obj_tag(x_20) == 0) +x_18 = l_Lean_Elab_Term_mkConst(x_16, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_18) == 0) { -lean_object* x_21; lean_object* x_22; -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -lean_ctor_set(x_13, 0, x_21); +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_13, 0, x_19); lean_ctor_set(x_3, 1, x_2); x_2 = x_3; x_3 = x_15; -x_10 = x_22; +x_10 = x_20; goto _start; } else { -uint8_t x_24; +uint8_t x_22; lean_free_object(x_13); lean_dec(x_17); lean_free_object(x_3); @@ -53898,71 +53926,66 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_24 = !lean_is_exclusive(x_20); -if (x_24 == 0) +x_22 = !lean_is_exclusive(x_18); +if (x_22 == 0) { -return x_20; +return x_18; } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; -x_25 = lean_ctor_get(x_20, 0); -x_26 = lean_ctor_get(x_20, 1); -lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_20); -x_27 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_27, 0, x_25); -lean_ctor_set(x_27, 1, x_26); -return x_27; +lean_object* x_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_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; -x_28 = lean_ctor_get(x_3, 1); -x_29 = lean_ctor_get(x_13, 0); -x_30 = lean_ctor_get(x_13, 1); -lean_inc(x_30); -lean_inc(x_29); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; +x_26 = lean_ctor_get(x_3, 1); +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_dec(x_13); -lean_inc(x_29); -x_31 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_29, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_32 = lean_ctor_get(x_31, 1); -lean_inc(x_32); -lean_dec(x_31); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_4); lean_inc(x_1); -x_33 = l_Lean_Elab_Term_mkConst(x_29, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_32); -if (lean_obj_tag(x_33) == 0) +x_29 = l_Lean_Elab_Term_mkConst(x_27, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_29) == 0) { -lean_object* x_34; lean_object* x_35; lean_object* x_36; -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_alloc_ctor(0, 2, 0); -lean_ctor_set(x_36, 0, x_34); -lean_ctor_set(x_36, 1, x_30); +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 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_30); +lean_ctor_set(x_32, 1, x_28); lean_ctor_set(x_3, 1, x_2); -lean_ctor_set(x_3, 0, x_36); +lean_ctor_set(x_3, 0, x_32); x_2 = x_3; -x_3 = x_28; -x_10 = x_35; +x_3 = x_26; +x_10 = x_31; goto _start; } else { -lean_object* x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; -lean_dec(x_30); -lean_free_object(x_3); +lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_dec(x_28); +lean_free_object(x_3); +lean_dec(x_26); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -53970,90 +53993,85 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_38 = lean_ctor_get(x_33, 0); -lean_inc(x_38); -x_39 = lean_ctor_get(x_33, 1); -lean_inc(x_39); -if (lean_is_exclusive(x_33)) { - lean_ctor_release(x_33, 0); - lean_ctor_release(x_33, 1); - x_40 = x_33; +x_34 = lean_ctor_get(x_29, 0); +lean_inc(x_34); +x_35 = lean_ctor_get(x_29, 1); +lean_inc(x_35); +if (lean_is_exclusive(x_29)) { + lean_ctor_release(x_29, 0); + lean_ctor_release(x_29, 1); + x_36 = x_29; } else { - lean_dec_ref(x_33); - x_40 = lean_box(0); + lean_dec_ref(x_29); + x_36 = lean_box(0); } -if (lean_is_scalar(x_40)) { - x_41 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_36)) { + x_37 = lean_alloc_ctor(1, 2, 0); } else { - x_41 = x_40; + x_37 = x_36; } -lean_ctor_set(x_41, 0, x_38); -lean_ctor_set(x_41, 1, x_39); -return x_41; +lean_ctor_set(x_37, 0, x_34); +lean_ctor_set(x_37, 1, x_35); +return x_37; } } } else { -lean_object* x_42; lean_object* x_43; lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; -x_42 = lean_ctor_get(x_3, 0); -x_43 = lean_ctor_get(x_3, 1); -lean_inc(x_43); -lean_inc(x_42); +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_3, 0); +x_39 = lean_ctor_get(x_3, 1); +lean_inc(x_39); +lean_inc(x_38); lean_dec(x_3); -x_44 = lean_ctor_get(x_42, 0); -lean_inc(x_44); -x_45 = lean_ctor_get(x_42, 1); -lean_inc(x_45); -if (lean_is_exclusive(x_42)) { - lean_ctor_release(x_42, 0); - lean_ctor_release(x_42, 1); - x_46 = x_42; +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_42); - x_46 = lean_box(0); + lean_dec_ref(x_38); + x_42 = lean_box(0); } -lean_inc(x_44); -x_47 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_44, x_4, x_5, x_6, x_7, x_8, x_9, x_10); -x_48 = lean_ctor_get(x_47, 1); -lean_inc(x_48); -lean_dec(x_47); lean_inc(x_9); lean_inc(x_8); lean_inc(x_7); lean_inc(x_6); lean_inc(x_4); lean_inc(x_1); -x_49 = l_Lean_Elab_Term_mkConst(x_44, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_48); -if (lean_obj_tag(x_49) == 0) +x_43 = l_Lean_Elab_Term_mkConst(x_40, x_1, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +if (lean_obj_tag(x_43) == 0) { -lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; -x_50 = lean_ctor_get(x_49, 0); -lean_inc(x_50); -x_51 = lean_ctor_get(x_49, 1); -lean_inc(x_51); -lean_dec(x_49); -if (lean_is_scalar(x_46)) { - x_52 = lean_alloc_ctor(0, 2, 0); +lean_object* x_44; lean_object* x_45; lean_object* x_46; lean_object* x_47; +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); +if (lean_is_scalar(x_42)) { + x_46 = lean_alloc_ctor(0, 2, 0); } else { - x_52 = x_46; + x_46 = x_42; } -lean_ctor_set(x_52, 0, x_50); -lean_ctor_set(x_52, 1, x_45); -x_53 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_53, 0, x_52); -lean_ctor_set(x_53, 1, x_2); -x_2 = x_53; -x_3 = x_43; -x_10 = x_51; +lean_ctor_set(x_46, 0, x_44); +lean_ctor_set(x_46, 1, x_41); +x_47 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_47, 0, x_46); +lean_ctor_set(x_47, 1, x_2); +x_2 = x_47; +x_3 = x_39; +x_10 = x_45; goto _start; } else { -lean_object* x_55; lean_object* x_56; lean_object* x_57; lean_object* x_58; -lean_dec(x_46); -lean_dec(x_45); -lean_dec(x_43); +lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; +lean_dec(x_42); +lean_dec(x_41); +lean_dec(x_39); lean_dec(x_9); lean_dec(x_8); lean_dec(x_7); @@ -54061,26 +54079,26 @@ lean_dec(x_6); lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_55 = lean_ctor_get(x_49, 0); -lean_inc(x_55); -x_56 = lean_ctor_get(x_49, 1); -lean_inc(x_56); -if (lean_is_exclusive(x_49)) { - lean_ctor_release(x_49, 0); - lean_ctor_release(x_49, 1); - x_57 = x_49; +x_49 = lean_ctor_get(x_43, 0); +lean_inc(x_49); +x_50 = lean_ctor_get(x_43, 1); +lean_inc(x_50); +if (lean_is_exclusive(x_43)) { + lean_ctor_release(x_43, 0); + lean_ctor_release(x_43, 1); + x_51 = x_43; } else { - lean_dec_ref(x_49); - x_57 = lean_box(0); + lean_dec_ref(x_43); + x_51 = lean_box(0); } -if (lean_is_scalar(x_57)) { - x_58 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_51)) { + x_52 = lean_alloc_ctor(1, 2, 0); } else { - x_58 = x_57; + x_52 = x_51; } -lean_ctor_set(x_58, 0, x_55); -lean_ctor_set(x_58, 1, x_56); -return x_58; +lean_ctor_set(x_52, 0, x_49); +lean_ctor_set(x_52, 1, x_50); +return x_52; } } } @@ -54091,44 +54109,16 @@ LEAN_EXPORT lean_object* l___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts(l { lean_object* x_10; lean_object* x_11; x_10 = lean_box(0); -x_11 = l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__3(x_2, x_10, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); +x_11 = l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_2, x_10, x_1, x_3, x_4, x_5, x_6, x_7, x_8, x_9); lean_dec(x_4); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_logWarning___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__2(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { -_start: -{ -lean_object* x_9; -x_9 = l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); -lean_dec(x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_9; -} -} -LEAN_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_EXPORT lean_object* l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___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_11; -x_11 = l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); +x_11 = l_List_foldlM___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8, x_9, x_10); lean_dec(x_5); return x_11; } @@ -60206,7 +60196,7 @@ LEAN_EXPORT lean_object* l_Lean_mkConstWithLevelParams___at_Lean_Elab_Term_expan { lean_object* x_9; lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -64256,32 +64246,32 @@ l_Lean_Elab_Term_isLetRecAuxMVar___closed__3 = _init_l_Lean_Elab_Term_isLetRecAu lean_mark_persistent(l_Lean_Elab_Term_isLetRecAuxMVar___closed__3); l_Lean_Elab_Term_isLetRecAuxMVar___closed__4 = _init_l_Lean_Elab_Term_isLetRecAuxMVar___closed__4(); lean_mark_persistent(l_Lean_Elab_Term_isLetRecAuxMVar___closed__4); -l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1 = _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1(); -lean_mark_persistent(l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__1); -l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2 = _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2(); -lean_mark_persistent(l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1___closed__2); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__1); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__2); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__3); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__4); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__5); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__6); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__7); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__8); +l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9 = _init_l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9(); +lean_mark_persistent(l_Lean_Linter_checkDeprecated___at_Lean_Elab_Term_mkConst___spec__1___closed__9); +l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1 = _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1(); +lean_mark_persistent(l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__1); +l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2 = _init_l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2(); +lean_mark_persistent(l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3___closed__2); l_Lean_Elab_Term_mkConst___closed__1 = _init_l_Lean_Elab_Term_mkConst___closed__1(); lean_mark_persistent(l_Lean_Elab_Term_mkConst___closed__1); l_Lean_Elab_Term_mkConst___closed__2 = _init_l_Lean_Elab_Term_mkConst___closed__2(); lean_mark_persistent(l_Lean_Elab_Term_mkConst___closed__2); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__1); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__2); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__3); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__4); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__5); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__6); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__7); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__8); -l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9 = _init_l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9(); -lean_mark_persistent(l_Lean_Linter_checkDeprecated___at___private_Lean_Elab_Term_0__Lean_Elab_Term_mkConsts___spec__1___closed__9); l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__1 = _init_l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__1(); lean_mark_persistent(l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__1); l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__2 = _init_l_Lean_Elab_throwAutoBoundImplicitLocal___at_Lean_Elab_Term_resolveName_process___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/Elab/Util.c b/stage0/stdlib/Lean/Elab/Util.c index 523feccc2c3b..576fc8564023 100644 --- a/stage0/stdlib/Lean/Elab/Util.c +++ b/stage0/stdlib/Lean/Elab/Util.c @@ -24,13 +24,14 @@ static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__22; lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1(lean_object*); static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__4; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__6; lean_object* l_Lean_Attribute_Builtin_getId(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_instMonadMacroAdapterOfMonadLift(lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandMacroImpl_x3f___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13; static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__3; extern lean_object* l_Lean_maxRecDepthErrorMessage; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__6; @@ -40,9 +41,7 @@ lean_object* l_Lean_indentD(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__2; lean_object* l_Lean_Elab_throwUnsupportedSyntax___rarg(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__18; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2; lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__8; LEAN_EXPORT lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__7; @@ -52,11 +51,12 @@ lean_object* l_Lean_Name_toString(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces___at_Lean_Elab_checkSyntaxNodeKindAtCurrentNamespaces___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind(lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__14; static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__4; lean_object* l_Lean_getRefPos___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10; static lean_object* l_Lean_Elab_mkElabAttribute___rarg___closed__2; lean_object* l_String_toFormat(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18; static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1; lean_object* l_List_foldl___at_Lean_MacroScopesView_review___spec__1(lean_object*, lean_object*); lean_object* l_Lean_declareBuiltinDocStringAndRanges(lean_object*, lean_object*, lean_object*, lean_object*); @@ -66,10 +66,10 @@ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Elab_adaptMacro___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_liftMacroM___rarg___lambda__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_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___rarg(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__9; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7; lean_object* l_List_find_x3f___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors(lean_object*, lean_object*); @@ -98,7 +98,9 @@ static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__24; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_getBetterRef___lambda__1___boxed(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15; LEAN_EXPORT lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__9(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_liftExcept___at_Lean_Elab_liftMacroM___spec__1(lean_object*, lean_object*, lean_object*); @@ -110,7 +112,6 @@ lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lea LEAN_EXPORT lean_object* l_Lean_MacroScopesView_format(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofSyntax(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5; static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__2; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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*, lean_object*, lean_object*, lean_object*); @@ -122,7 +123,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNes static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__16; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Elab_addMacroStack___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Elab_getBetterRef___lambda__1(lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); @@ -133,7 +133,6 @@ LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg_ static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__5; static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_withLogging___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__4; static lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam___closed__2; static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__1; lean_object* lean_st_ref_get(lean_object*, lean_object*); @@ -145,14 +144,15 @@ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__1(lean_object*, LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_syntaxNodeKindOfAttrParam___closed__3; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__6; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Macro_hasDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_logException___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1___closed__4; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_expandMacroImpl_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -161,13 +161,13 @@ LEAN_EXPORT lean_object* l_Lean_Elab_logException(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_macroAttribute; LEAN_EXPORT lean_object* l_Lean_Elab_pp_macroStack; static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__7; +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004_(lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__5; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11; LEAN_EXPORT lean_object* l_Lean_Elab_getBetterRef___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12; static lean_object* l_Lean_Elab_checkSyntaxNodeKind___rarg___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__10; @@ -182,16 +182,15 @@ static lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1_ uint8_t l_Lean_Environment_contains(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_ReaderT_pure___at_Lean_Elab_liftMacroM___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandOptNamedPrio___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__2; static lean_object* l___private_Lean_Elab_Util_0__Lean_Elab_evalSyntaxConstantUnsafe___closed__2; static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__4; static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__3; lean_object* l_Lean_Syntax_reprint(lean_object*); lean_object* lean_name_append_index_after(lean_object*, lean_object*); static lean_object* l_Lean_Elab_expandMacroImpl_x3f___closed__1; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__26; -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034_(lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__15; lean_object* l_Lean_throwErrorAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandOptNamedPrio(lean_object*, lean_object*, lean_object*); @@ -208,28 +207,25 @@ static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__12; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Elab_mkMacroAttributeUnsafe___closed__1; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16; LEAN_EXPORT lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_mkElabAttribute___rarg___lambda__3(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_environment_main_module(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__3; LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___at_Lean_Elab_checkSyntaxNodeKindAtCurrentNamespaces___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__3; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__13; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Elab_syntaxNodeKindOfAttrParam___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__14; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__1; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__25; LEAN_EXPORT lean_object* l_Lean_Elab_adaptMacro(lean_object*); static lean_object* l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2; lean_object* l_Lean_isTracingEnabledFor___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_logAt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__12; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Syntax_isNone(lean_object*); LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3(lean_object*); @@ -243,11 +239,14 @@ lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object* lean_object* l_Lean_toMessageList(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_expandMacroImpl_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_addMacroStack___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4; static lean_object* l_Lean_Elab_logException___rarg___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKind___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_prettyPrint(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___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*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9; static lean_object* l_Lean_Elab_expandOptNamedPrio___closed__1; static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__23; size_t lean_usize_add(size_t, size_t); @@ -268,13 +267,14 @@ lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_o lean_object* l_Lean_Syntax_unsetTrailing(lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); static lean_object* l_Lean_Elab_nestedExceptionToMessageData___rarg___lambda__1___closed__2; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1; uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__6; LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___rarg___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17; static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__21; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Elab_liftMacroM___spec__3___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8; +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2; static lean_object* l_Lean_Elab_addMacroStack___rarg___lambda__1___closed__1; uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_nat_add(lean_object*, lean_object*); @@ -291,13 +291,13 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_withLogging(lean_object*); LEAN_EXPORT lean_object* l_Lean_Elab_checkSyntaxNodeKindAtNamespaces___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1; uint8_t l_List_beq___at_Lean_MacroScopesView_isPrefixOf___spec__1(lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17; LEAN_EXPORT lean_object* l_Lean_MacroScopesView_equalScope___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Elab_Util___hyg_777____closed__11; LEAN_EXPORT 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_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7; -static lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__10; LEAN_EXPORT uint8_t l_Lean_MacroScopesView_equalScope(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Syntax_prettyPrint(lean_object* x_1) { _start: @@ -2518,7 +2518,7 @@ x_8 = l_Lean_Elab_mkElabAttribute___rarg(x_3, x_4, x_5, x_6, x_7, x_1, x_2); return x_8; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1() { _start: { lean_object* x_1; @@ -2526,22 +2526,22 @@ x_1 = lean_mk_string_from_bytes("macroAttribute", 14); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_expandOptNamedPrio___closed__1; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__6; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2; x_3 = l_Lean_Elab_mkMacroAttributeUnsafe(x_2, x_1); return x_3; } @@ -4650,7 +4650,7 @@ x_8 = l_Array_mapMUnsafe_map___at_Lean_Elab_throwErrorWithNestedErrors___spec__1 return x_8; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4660,7 +4660,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_Util___hyg_2533____closed__2() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -4670,17 +4670,17 @@ 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_Util___hyg_2533____closed__3() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__2; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____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_Util___hyg_2533____closed__4() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4() { _start: { lean_object* x_1; @@ -4688,17 +4688,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__3; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__4; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4; 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_Util___hyg_2533____closed__6() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__6() { _start: { lean_object* x_1; @@ -4706,37 +4706,37 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__6; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____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_Util___hyg_2533____closed__8() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7; x_2 = l_Lean_Elab_expandOptNamedPrio___closed__1; 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_Util___hyg_2533____closed__9() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__8; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8; x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____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_Util___hyg_2533____closed__10() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10() { _start: { lean_object* x_1; @@ -4744,17 +4744,17 @@ x_1 = lean_mk_string_from_bytes("Util", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__9; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__10; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10; 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_Util___hyg_2533____closed__12() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12() { _start: { lean_object* x_1; @@ -4762,27 +4762,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__12; +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12; 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_Util___hyg_2533____closed__14() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13; -x_2 = lean_unsigned_to_nat(2533u); +x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13; +x_2 = lean_unsigned_to_nat(2503u); 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_Util___hyg_2533____closed__15() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15() { _start: { lean_object* x_1; @@ -4790,17 +4790,17 @@ x_1 = lean_mk_string_from_bytes("step", 4); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__6; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15; x_3 = l_Lean_Name_mkStr2(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17() { _start: { lean_object* x_1; @@ -4808,24 +4808,24 @@ x_1 = lean_mk_string_from_bytes("result", 6); return x_1; } } -static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18() { +static lean_object* _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_341____closed__6; -x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15; -x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15; +x_3 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17; x_4 = l_Lean_Name_mkStr3(x_1, x_2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503_(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_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1; +x_2 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1; x_3 = 0; -x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__14; +x_4 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -4833,7 +4833,7 @@ lean_object* x_6; lean_object* x_7; lean_object* x_8; x_6 = lean_ctor_get(x_5, 1); lean_inc(x_6); lean_dec(x_5); -x_7 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16; +x_7 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16; x_8 = l_Lean_registerTraceClass(x_7, x_3, x_4, x_6); if (lean_obj_tag(x_8) == 0) { @@ -4841,7 +4841,7 @@ lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); -x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18; +x_10 = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18; x_11 = 1; x_12 = l_Lean_registerTraceClass(x_10, x_11, x_4, x_9); return x_12; @@ -5045,11 +5045,11 @@ l_Lean_Elab_mkMacroAttributeUnsafe___closed__5 = _init_l_Lean_Elab_mkMacroAttrib lean_mark_persistent(l_Lean_Elab_mkMacroAttributeUnsafe___closed__5); l_Lean_Elab_mkMacroAttributeUnsafe___closed__6 = _init_l_Lean_Elab_mkMacroAttributeUnsafe___closed__6(); lean_mark_persistent(l_Lean_Elab_mkMacroAttributeUnsafe___closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034____closed__2); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1034_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004____closed__2); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_1004_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Elab_macroAttribute = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Elab_macroAttribute); @@ -5078,43 +5078,43 @@ l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1 = _init_l_ lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__1); l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2 = _init_l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2(); lean_mark_persistent(l_Lean_Elab_throwErrorWithNestedErrors___rarg___lambda__1___closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__1); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__2(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__2); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__3(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__3); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__4(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__4); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__5); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__6(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__6); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__7); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__8(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__8); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__9(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__9); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__10(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__10); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__11); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__12 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__12(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__12); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__13); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__14 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__14(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__14); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__15); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__16); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__17); -l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18(); -lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533____closed__18); -if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2533_(lean_io_mk_world()); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__1); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__2); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__3); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__4); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__5); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__6 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__6(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__6); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__7); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__8); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__9); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__10); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__11); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__12); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__13); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__14); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__15); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__16); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__17); +l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18 = _init_l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18(); +lean_mark_persistent(l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503____closed__18); +if (builtin) {res = l_Lean_Elab_initFn____x40_Lean_Elab_Util___hyg_2503_(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/Environment.c b/stage0/stdlib/Lean/Environment.c index 3c454cfcc015..acf67902f15f 100644 --- a/stage0/stdlib/Lean/Environment.c +++ b/stage0/stdlib/Lean/Environment.c @@ -13,6 +13,7 @@ #ifdef __cplusplus extern "C" { #endif +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__15; LEAN_EXPORT uint32_t l_Lean_EnvironmentHeader_trustLevel___default; uint8_t lean_compacted_region_is_memory_mapped(size_t); static lean_object* l_Lean_TagDeclarationExtension_instInhabited___closed__1; @@ -21,12 +22,15 @@ LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_obj 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*); LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__17; 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_Lean_Name_reprPrec(lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_mkModuleData___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__2; static lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg___closed__2; @@ -40,13 +44,13 @@ static lean_object* l_Lean_Kernel_enableDiag___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkModuleData(lean_object*, lean_object*); lean_object* l_Lean_Name_quickLt___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2(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); 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*); +lean_object* lean_add_decl(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__7(lean_object*, lean_object*); static lean_object* l_Lean_Environment_displayStats___closed__10; @@ -54,10 +58,11 @@ LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe(lean_object LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries(lean_object*, lean_object*); lean_object* lean_uint32_to_nat(uint32_t); static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__3; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__18; 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___auto____x40_Lean_Environment___hyg_2510____closed__9; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__27; 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*); @@ -71,7 +76,6 @@ static lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpe LEAN_EXPORT lean_object* l_Lean_importModules___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__7(lean_object*, size_t, size_t, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__15; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_modifyState___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag___boxed(lean_object*, lean_object*); @@ -85,14 +89,11 @@ LEAN_EXPORT lean_object* l_Lean_EnvironmentHeader_regions___default; 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; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__26; 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___auto____x40_Lean_Environment___hyg_2509____closed__28; LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__2___boxed(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__6; 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*); size_t lean_uint64_to_usize(uint64_t); @@ -101,7 +102,6 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Environme 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*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_mkMapDeclarationExtension___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_registerPersistentEnvExtensionUnsafe___spec__1(lean_object*, lean_object*, lean_object*); @@ -116,6 +116,7 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_mkExtNameMap___spec__3(l static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___closed__1; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_finalizeImport___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__21; 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*); @@ -141,7 +142,6 @@ extern uint8_t l_instInhabitedBool; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3(lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg___closed__1; static lean_object* l_Lean_mkTagDeclarationExtension___closed__3; -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__10(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_readModuleData___boxed(lean_object*, lean_object*); @@ -162,26 +162,26 @@ LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterf LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__1(lean_object*, lean_object*); 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*); +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1381_(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__1; lean_object* l_Lean_RBNode_insert___at_Lean_NameSet_insert___spec__1(lean_object*, lean_object*, lean_object*); 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*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_imports___boxed(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__23; 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_2510____closed__23; 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*); -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* lean_kernel_get_diag(lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportStateM_run(lean_object*); @@ -191,15 +191,12 @@ LEAN_EXPORT lean_object* l_Lean_HashMap_insertIfNew___at_Lean_finalizeImport___s LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_instInhabited___rarg(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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__5; LEAN_EXPORT lean_object* l_panic___at_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___lambda__1(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__4; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_mkInitialExtensionStates(lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___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*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__2; 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*); @@ -209,6 +206,7 @@ LEAN_EXPORT uint8_t l_Lean_MapDeclarationExtension_contains___rarg(lean_object*, LEAN_EXPORT lean_object* l_Lean_EnvExtensionEntrySpec; LEAN_EXPORT lean_object* l_Lean_updateEnvAttributes___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__3(lean_object*, size_t, size_t, 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_object* l_System_FilePath_pathExists(lean_object*, lean_object*); @@ -232,13 +230,9 @@ LEAN_EXPORT lean_object* l_Lean_Kernel_enableDiag(lean_object*, uint8_t); 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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__24; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_finalizeImport___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Kernel_isDiagnosticsEnabled(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__14; LEAN_EXPORT lean_object* l_panic___at_Lean_throwAlreadyImported___spec__1(lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6541_(lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__5(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_registerSimplePersistentEnvExtension(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_finalizeImport___spec__8___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -246,14 +240,12 @@ static lean_object* l_Lean_TagDeclarationExtension_tag___closed__1; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__4(lean_object*); lean_object* l_IO_println___at_Lean_instEval___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__6; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtensionDescr_toArrayFn___default(lean_object*); 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*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkMapDeclarationExtension___spec__1___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__16; 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*); size_t lean_usize_of_nat(lean_object*); @@ -264,17 +256,21 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_mkExtNameMap___spec 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*); LEAN_EXPORT lean_object* l_Lean_Kernel_Diagnostics_isEnabled___boxed(lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__1; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__6; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_displayStats___spec__8___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_SMap_numBuckets___at_Lean_Environment_displayStats___spec__5(lean_object*); lean_object* l_panic___rarg(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3340_; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__2; static lean_object* l_Lean_Environment_displayStats___closed__3; static lean_object* l_Lean_Kernel_instInhabitedDiagnostics___closed__1; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_mkModuleData___spec__4(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__2; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__3; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__10; LEAN_EXPORT lean_object* l_Lean_instCoeNameImport(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__11; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__11; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__9; lean_object* lean_st_ref_take(lean_object*, lean_object*); @@ -286,6 +282,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState(lean_object*, le 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*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298_(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*); @@ -295,7 +292,9 @@ LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__2(lean_object*, lean_ob LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__2; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__28; 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*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType___rarg(lean_object*, lean_object*); @@ -310,14 +309,13 @@ LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_finalizeImport___spec__2 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*); LEAN_EXPORT lean_object* l_Lean_instModuleIdxBEq; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3192_; lean_object* l_Lean_HashSetImp_insert___at_Lean_NameHashSet_insert___spec__1(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3251_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Kernel_Diagnostics_recordUnfold___spec__4(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkEmptyEnvironment___lambda__1(uint32_t, lean_object*, lean_object*); static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__4; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_3429_; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); @@ -333,24 +331,24 @@ LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift___rarg___lambda__1(lean_ static lean_object* l_Lean_TagDeclarationExtension_tag___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_getTrustLevel___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_importModules___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__25; LEAN_EXPORT lean_object* l_Lean_ImportState_regions___default; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_isDefEq___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_display_stats(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__13; 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*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__1; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__18; 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; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__17; static lean_object* l_Lean_throwAlreadyImported___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_find_x3f(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4; LEAN_EXPORT uint8_t l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1___rarg(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__20; lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__3___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_registerNamePrefixes(lean_object*, lean_object*); @@ -369,14 +367,14 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__2___r lean_object* lean_array_to_list(lean_object*, lean_object*); static lean_object* l_Lean_instInhabitedModuleData___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_ensureExtensionsArraySize(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2912_; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_mkExtNameMap___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_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_object* l_Lean_Expr_FindImpl_findUnsafeM_x3f___at_Lean_Expr_FindImpl_findUnsafe_x3f___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1; LEAN_EXPORT lean_object* l_Lean_withImportModules___rarg(lean_object*, lean_object*, uint32_t, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__3; lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__4; 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; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_MapDeclarationExtension_contains___spec__1(lean_object*); @@ -384,6 +382,8 @@ 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; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__5(lean_object*, size_t, size_t, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Environment_addDeclCore___boxed(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*); LEAN_EXPORT lean_object* l_Lean_RBTree_toArray___at_Lean_mkModuleData___spec__7(lean_object*); LEAN_EXPORT lean_object* l_Lean_ImportState_moduleNameSet___default; @@ -393,7 +393,6 @@ LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Environment_addE 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; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__12; 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_object* l_List_lengthTRAux___rarg(lean_object*, lean_object*); @@ -430,6 +429,7 @@ static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed_ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__26; 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*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2468_(lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__14; LEAN_EXPORT lean_object* l_panic___at_Lean_MapDeclarationExtension_insert___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_instToStringImport___closed__2; @@ -443,11 +443,13 @@ LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtension___lambda__4( 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*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6422_(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withImportModules(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__12; 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___auto____x40_Lean_Environment___hyg_2509____closed__11; static lean_object* l_Lean_mkEmptyEnvironment___closed__2; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3(lean_object*); extern lean_object* l_Std_Format_defWidth; @@ -457,7 +459,6 @@ LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries___rarg(l LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Environment_freeRegions___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_MapDeclarationExtension_instInhabited___closed__1; LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1___lambda__1(lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, lean_object*); 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; @@ -465,19 +466,19 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_getModu 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___auto____x40_Lean_Environment___hyg_2510_; 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*); 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_Array_binSearchAux___at_Lean_TagDeclarationExtension_isTagged___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__7; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModules___spec__1(lean_object*, size_t, size_t, lean_object*, lean_object*); 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; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2___closed__1; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5; LEAN_EXPORT lean_object* l_Lean_throwAlreadyImported___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkEmptyEnvironment___spec__2(lean_object*); static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__1; @@ -487,6 +488,7 @@ LEAN_EXPORT lean_object* l_Lean_SMap_contains___at_Lean_Environment_addExtraName LEAN_EXPORT lean_object* l_Lean_finalizeImport___lambda__1___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_HashMapImp_contains___at_Lean_Environment_addExtraName___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_throwUnexpectedType(lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState(lean_object*); LEAN_EXPORT lean_object* l_List_toString___at_Lean_Environment_displayStats___spec__1(lean_object*); static lean_object* l_Lean_Environment_registerNamespace___closed__1; @@ -495,12 +497,15 @@ LEAN_EXPORT lean_object* l_Lean_Kernel_Diagnostics_unfoldCounter___default; 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; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__10; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__2___boxed(lean_object*, lean_object*, 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*); LEAN_EXPORT lean_object* l_Lean_Kernel_instInhabitedDiagnostics; LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_EnvExtensionInterfaceUnsafe_invalidExtMsg; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp; +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2(lean_object*); 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_Environment_hasUnsafe___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getEntries(lean_object*, lean_object*); @@ -509,10 +514,8 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_setState___rarg___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__8; LEAN_EXPORT lean_object* l_Lean_Environment_hasUnsafe___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___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___auto____x40_Lean_Environment___hyg_2509_; static lean_object* l_Lean_instInhabitedEnvExtensionInterface___closed__5; static lean_object* l_Lean_MapDeclarationExtension_insert___rarg___closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg___lambda__1(lean_object*, lean_object*); @@ -520,14 +523,10 @@ static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_ 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 uint8_t lean_kernel_diag_is_enabled(lean_object*); -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1; -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__2(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*); LEAN_EXPORT lean_object* lean_environment_mark_quot_init(lean_object*); -LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__4(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__7; LEAN_EXPORT lean_object* l_Lean_SMap_stageSizes___at_Lean_Environment_displayStats___spec__4___boxed(lean_object*); lean_object* lean_string_length(lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__6; @@ -549,6 +548,7 @@ 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; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5; 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*); @@ -568,7 +568,6 @@ lean_object* l_Lean_findOLean(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Import_runtimeOnly___default; LEAN_EXPORT lean_object* l_Lean_Environment_isConstructor___boxed(lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__27; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__20; LEAN_EXPORT lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__1(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_SMap_contains___at_Lean_Environment_addExtraName___spec__1(lean_object*, lean_object*); @@ -583,23 +582,26 @@ lean_object* lean_array_set(lean_object*, lean_object*, lean_object*); 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; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__20; 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_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*); extern lean_object* l_Id_instMonad; 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*); LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f_x27___at_Lean_Environment_find_x3f___spec__1(lean_object*, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__24; lean_object* lean_panic_fn(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_contains___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_finalizePersistentExtensions_loop___closed__1; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__26; 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_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6; 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*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Environment___hyg_2883_; LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface; static lean_object* l_Lean_instInhabitedPersistentEnvExtension___closed__5; lean_object* lean_nat_mul(lean_object*, lean_object*); @@ -607,12 +609,10 @@ LEAN_EXPORT uint8_t l_Lean_Kernel_isDefEqGuarded(lean_object*, lean_object*, lea static uint32_t l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___lambda__1___closed__1; 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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__2; LEAN_EXPORT lean_object* l_Lean_SimplePersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3f___spec__6___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Environment_addExtraName___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__9; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_addEntry(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__21; @@ -633,17 +633,13 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtension_instInhabited___boxed(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___auto____x40_Lean_Environment___hyg_2509____closed__1; -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1380_(lean_object*); LEAN_EXPORT lean_object* l_Lean_persistentEnvExtensionsRef; lean_object* lean_io_initializing(lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__3; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__1; LEAN_EXPORT lean_object* l_Std_Range_forIn_x27_loop___at_Lean_finalizeImport___spec__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_imports(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__21; static lean_object* l___private_Lean_Environment_0__Lean_reprImport____x40_Lean_Environment___hyg_75____closed__13; size_t lean_usize_sub(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_instInhabitedPersistentEnvExtensionState___rarg(lean_object*); @@ -658,6 +654,7 @@ LEAN_EXPORT lean_object* l_Lean_TagDeclarationExtension_isTagged___boxed(lean_ob 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___auto____x40_Lean_Environment___hyg_2510____closed__5; 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*); @@ -669,19 +666,18 @@ LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_modifyState___rarg(l LEAN_EXPORT lean_object* l_Lean_MapDeclarationExtension_instInhabited(lean_object*); extern lean_object* l_instBEqNat; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2467_(lean_object*); static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__2; static lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg___lambda__2___closed__1; extern lean_object* l_Lean_instInhabitedName; LEAN_EXPORT lean_object* l_Lean_Kernel_isDiagnosticsEnabled___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Environment_0__Lean_Environment_isQuotInit___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_instInhabitedEnvExtensionInterface___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417_(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_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; +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__22; LEAN_EXPORT lean_object* l_Lean_CompactedRegion_isMemoryMapped___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Kernel_resetDiag(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__1(lean_object*, lean_object*, lean_object*); @@ -701,6 +697,7 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Environment_find_x3 LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___elambda__3___rarg(lean_object*, lean_object*, lean_object*); 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___auto____x40_Lean_Environment___hyg_2510____closed__3; 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*); @@ -709,7 +706,6 @@ 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*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__19; static lean_object* l_Lean_mkEmptyEnvironment___lambda__1___closed__1; lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_evalConstCheck___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -719,19 +715,17 @@ static lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spe LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Environment_find_x3f___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_registerSimplePersistentEnvExtension___rarg___closed__1; +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__6; 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_2509____closed__13; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_mkExtNameMap___spec__1___boxed(lean_object*); -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__22; static lean_object* l_Lean_Environment_displayStats___closed__2; -static lean_object* l___auto____x40_Lean_Environment___hyg_2509____closed__25; -static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__3; LEAN_EXPORT lean_object* l_Lean_saveModuleData___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_instMonadEnvOfMonadLift___rarg(lean_object*, 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___auto____x40_Lean_Environment___hyg_2510____closed__16; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Kernel_Diagnostics_recordUnfold___spec__7(lean_object*, 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*); @@ -749,9 +743,11 @@ 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_Kernel_enableDiag___lambda__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkTagDeclarationExtension(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1; 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* lean_import_modules(lean_object*, lean_object*, uint32_t, uint8_t, lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__19; 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; @@ -759,6 +755,7 @@ 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; LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_getState___rarg___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_TagDeclarationExtension_tag___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_EnvExtensionInterfaceUnsafe_mkInitialExtStates___spec__1(size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Subarray_forInUnsafe_loop___at___private_Lean_Environment_0__Lean_setImportedEntries___spec__1(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); @@ -774,6 +771,7 @@ lean_object* lean_array_uset(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* lean_environment_free_regions(lean_object*, lean_object*); lean_object* lean_array_get(lean_object*, lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__14; LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Environment_displayStats___spec__2(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_importModulesCore___spec__1___lambda__2___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Environment_addExtraName___spec__3___boxed(lean_object*, lean_object*); @@ -788,11 +786,13 @@ 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*); +static lean_object* l___auto____x40_Lean_Environment___hyg_2510____closed__7; LEAN_EXPORT lean_object* l_Lean_PersistentEnvExtension_setState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__2(lean_object*, size_t, size_t, lean_object*); lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); +static lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1; 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*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__3(lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_mkTagDeclarationExtension___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Environment_0__Lean_Environment_addAux___spec__3___closed__3; static lean_object* l_Lean_EnvExtensionInterfaceUnsafe_imp___closed__3; @@ -3591,13 +3591,15 @@ lean_dec(x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Environment_addDecl___boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Environment_addDeclCore___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_3; -x_3 = lean_add_decl(x_1, x_2); +size_t x_4; lean_object* x_5; +x_4 = lean_unbox_usize(x_2); lean_dec(x_2); -return x_3; +x_5 = lean_add_decl(x_1, x_4, x_3); +lean_dec(x_3); +return x_5; } } LEAN_EXPORT lean_object* l_Lean_ConstantInfo_instantiateTypeLevelParams(lean_object* x_1, lean_object* x_2) { @@ -3812,7 +3814,7 @@ x_2 = l_Lean_EnvExtensionInterfaceUnsafe_instInhabitedExt___closed__2; return x_2; } } -LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1380_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1381_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -5633,7 +5635,7 @@ lean_dec(x_1); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2467_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_2468_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -5659,7 +5661,7 @@ return x_7; } } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__1() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__1() { _start: { lean_object* x_1; @@ -5667,7 +5669,7 @@ x_1 = lean_mk_string_from_bytes("Lean", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__2() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__2() { _start: { lean_object* x_1; @@ -5675,7 +5677,7 @@ x_1 = lean_mk_string_from_bytes("Parser", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__3() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__3() { _start: { lean_object* x_1; @@ -5683,7 +5685,7 @@ x_1 = lean_mk_string_from_bytes("Tactic", 6); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__4() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__4() { _start: { lean_object* x_1; @@ -5691,19 +5693,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq", 9); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__5() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2509____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2509____closed__4; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__6() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__6() { _start: { lean_object* x_1; @@ -5711,19 +5713,19 @@ x_1 = lean_mk_string_from_bytes("tacticSeq1Indented", 18); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__7() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2509____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2509____closed__6; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__8() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__8() { _start: { lean_object* x_1; @@ -5731,17 +5733,17 @@ x_1 = lean_mk_string_from_bytes("null", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__9() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__8; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__10() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__10() { _start: { lean_object* x_1; @@ -5749,41 +5751,41 @@ x_1 = lean_mk_string_from_bytes("exact", 5); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__11() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2509____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__3; -x_4 = l___auto____x40_Lean_Environment___hyg_2509____closed__10; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____closed__3; +x_4 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__12() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__10; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__13() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__12; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__12; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__14() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__14() { _start: { lean_object* x_1; @@ -5791,7 +5793,7 @@ x_1 = lean_mk_string_from_bytes("Term", 4); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__15() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__15() { _start: { lean_object* x_1; @@ -5799,19 +5801,19 @@ x_1 = lean_mk_string_from_bytes("declName", 8); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__16() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__1; -x_2 = l___auto____x40_Lean_Environment___hyg_2509____closed__2; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__14; -x_4 = l___auto____x40_Lean_Environment___hyg_2509____closed__15; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__1; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__2; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____closed__14; +x_4 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__17() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__17() { _start: { lean_object* x_1; @@ -5819,35 +5821,35 @@ x_1 = lean_mk_string_from_bytes("decl_name%", 10); return x_1; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__18() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__17; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__19() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__18; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__18; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__20() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__16; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__19; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__16; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____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); @@ -5855,23 +5857,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__21() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____closed__21() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__13; -x_2 = l___auto____x40_Lean_Environment___hyg_2509____closed__20; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__13; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__20; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__22() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__11; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__21; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__11; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____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); @@ -5879,23 +5881,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__23() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__22; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__22; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__24() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__9; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__23; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__9; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____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); @@ -5903,23 +5905,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__25() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__24; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__24; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__26() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__7; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__25; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__7; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____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); @@ -5927,23 +5929,23 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__27() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__26; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__26; x_3 = lean_array_push(x_1, x_2); return x_3; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509____closed__28() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510____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_2509____closed__5; -x_3 = l___auto____x40_Lean_Environment___hyg_2509____closed__27; +x_2 = l___auto____x40_Lean_Environment___hyg_2510____closed__5; +x_3 = l___auto____x40_Lean_Environment___hyg_2510____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); @@ -5951,11 +5953,11 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2509_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2510_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__28; return x_1; } } @@ -6413,11 +6415,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_2912_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_2883_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__28; return x_1; } } @@ -6825,11 +6827,11 @@ lean_dec(x_1); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3251_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3192_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__28; return x_1; } } @@ -7240,11 +7242,11 @@ x_5 = lean_box(x_4); return x_5; } } -static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3429_() { +static lean_object* _init_l___auto____x40_Lean_Environment___hyg_3340_() { _start: { lean_object* x_1; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__28; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__28; return x_1; } } @@ -13890,7 +13892,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_6417____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_6298____spec__2(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13913,7 +13915,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____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_6298____spec__3(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -13951,7 +13953,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_6417____spec__2(x_6, x_15, x_16, x_4); +x_17 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__2(x_6, x_15, x_16, x_4); lean_dec(x_6); x_2 = x_11; x_4 = x_17; @@ -13965,7 +13967,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -13994,14 +13996,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_6417____spec__3(x_2, x_7, x_8, x_1); +x_9 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____spec__4(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(lean_object* x_1) { _start: { uint8_t x_2; @@ -14039,7 +14041,7 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6417____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_6298____spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4) { _start: { uint8_t x_5; @@ -14065,7 +14067,7 @@ return x_4; } } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; lean_object* x_4; @@ -14074,7 +14076,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_6417____lambda__2___closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -14083,7 +14085,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -14094,15 +14096,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_6417____lambda__2___closed__1; -x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__1(x_5, x_1); +x_5 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1; +x_6 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____spec__4(x_9); +x_10 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(x_9); return x_10; } else @@ -14113,15 +14115,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_6417____lambda__2___closed__1; -x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6417____spec__1(x_12, x_1); +x_12 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1; +x_13 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____spec__4(x_16); +x_17 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(x_16); return x_17; } else @@ -14130,23 +14132,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_6417____spec__5(x_1, x_18, x_19, x_3); +x_20 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____spec__1(x_21, x_1); +x_22 = l_Lean_mkStateFromImportedEntries___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____spec__4(x_25); +x_26 = l_Lean_SMap_switch___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__4(x_25); return x_26; } } } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__1() { _start: { lean_object* x_1; @@ -14154,17 +14156,17 @@ x_1 = lean_mk_string_from_bytes("namespacesExt", 13); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__2() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___auto____x40_Lean_Environment___hyg_2509____closed__1; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1; +x_1 = l___auto____x40_Lean_Environment___hyg_2510____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____closed__3() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__3() { _start: { lean_object* x_1; @@ -14173,30 +14175,30 @@ lean_closure_set(x_1, 0, lean_box(0)); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__1), 2, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2), 1, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2), 1, 0); return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____closed__2; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4; -x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5; -x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__3; +x_1 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4; +x_3 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5; +x_4 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____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); @@ -14205,16 +14207,16 @@ lean_ctor_set(x_5, 3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6417_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6298_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____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_6298____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; @@ -14222,12 +14224,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_6417____spec__2(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____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_6298____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; @@ -14235,12 +14237,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_6417____spec__3(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____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_6417____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_6298____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; @@ -14248,7 +14250,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_6417____spec__5(x_1, x_5, x_6, x_4); +x_7 = l_Array_foldlMUnsafe_fold___at_Lean_initFn____x40_Lean_Environment___hyg_6298____spec__5(x_1, x_5, x_6, x_4); lean_dec(x_1); return x_7; } @@ -14289,7 +14291,7 @@ x_1 = l_Lean_Kernel_instInhabitedDiagnostics___closed__1; return x_1; } } -static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1() { +static lean_object* _init_l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1() { _start: { lean_object* x_1; lean_object* x_2; @@ -14299,11 +14301,11 @@ lean_closure_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6541_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_Environment___hyg_6422_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1; +x_2 = l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1; x_3 = l_Lean_EnvExtensionInterfaceUnsafe_registerExt___rarg(x_2, x_1); return x_3; } @@ -17194,7 +17196,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_1380_(lean_io_mk_world()); +if (builtin) {res = l_Lean_EnvExtensionInterfaceUnsafe_initFn____x40_Lean_Environment___hyg_1381_(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); @@ -17261,69 +17263,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_2467_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_2468_(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_2509____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__1(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__1); -l___auto____x40_Lean_Environment___hyg_2509____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__2(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__2); -l___auto____x40_Lean_Environment___hyg_2509____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__3(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__3); -l___auto____x40_Lean_Environment___hyg_2509____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__4(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__4); -l___auto____x40_Lean_Environment___hyg_2509____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__5(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__5); -l___auto____x40_Lean_Environment___hyg_2509____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__6(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__6); -l___auto____x40_Lean_Environment___hyg_2509____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__7(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__7); -l___auto____x40_Lean_Environment___hyg_2509____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__8(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__8); -l___auto____x40_Lean_Environment___hyg_2509____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__9(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__9); -l___auto____x40_Lean_Environment___hyg_2509____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__10(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__10); -l___auto____x40_Lean_Environment___hyg_2509____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__11(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__11); -l___auto____x40_Lean_Environment___hyg_2509____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__12(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__12); -l___auto____x40_Lean_Environment___hyg_2509____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__13(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__13); -l___auto____x40_Lean_Environment___hyg_2509____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__14(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__14); -l___auto____x40_Lean_Environment___hyg_2509____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__15(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__15); -l___auto____x40_Lean_Environment___hyg_2509____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__16(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__16); -l___auto____x40_Lean_Environment___hyg_2509____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__17(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__17); -l___auto____x40_Lean_Environment___hyg_2509____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__18(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__18); -l___auto____x40_Lean_Environment___hyg_2509____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__19(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__19); -l___auto____x40_Lean_Environment___hyg_2509____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__20(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__20); -l___auto____x40_Lean_Environment___hyg_2509____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__21(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__21); -l___auto____x40_Lean_Environment___hyg_2509____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__22(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__22); -l___auto____x40_Lean_Environment___hyg_2509____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__23(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__23); -l___auto____x40_Lean_Environment___hyg_2509____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__24(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__24); -l___auto____x40_Lean_Environment___hyg_2509____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__25(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__25); -l___auto____x40_Lean_Environment___hyg_2509____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__26(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__26); -l___auto____x40_Lean_Environment___hyg_2509____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__27(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__27); -l___auto____x40_Lean_Environment___hyg_2509____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_2509____closed__28(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509____closed__28); -l___auto____x40_Lean_Environment___hyg_2509_ = _init_l___auto____x40_Lean_Environment___hyg_2509_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2509_); +}l___auto____x40_Lean_Environment___hyg_2510____closed__1 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__1(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__1); +l___auto____x40_Lean_Environment___hyg_2510____closed__2 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__2(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__2); +l___auto____x40_Lean_Environment___hyg_2510____closed__3 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__3(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__3); +l___auto____x40_Lean_Environment___hyg_2510____closed__4 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__4(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__4); +l___auto____x40_Lean_Environment___hyg_2510____closed__5 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__5(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__5); +l___auto____x40_Lean_Environment___hyg_2510____closed__6 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__6(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__6); +l___auto____x40_Lean_Environment___hyg_2510____closed__7 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__7(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__7); +l___auto____x40_Lean_Environment___hyg_2510____closed__8 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__8(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__8); +l___auto____x40_Lean_Environment___hyg_2510____closed__9 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__9(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__9); +l___auto____x40_Lean_Environment___hyg_2510____closed__10 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__10(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__10); +l___auto____x40_Lean_Environment___hyg_2510____closed__11 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__11(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__11); +l___auto____x40_Lean_Environment___hyg_2510____closed__12 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__12(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__12); +l___auto____x40_Lean_Environment___hyg_2510____closed__13 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__13(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__13); +l___auto____x40_Lean_Environment___hyg_2510____closed__14 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__14(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__14); +l___auto____x40_Lean_Environment___hyg_2510____closed__15 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__15(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__15); +l___auto____x40_Lean_Environment___hyg_2510____closed__16 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__16(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__16); +l___auto____x40_Lean_Environment___hyg_2510____closed__17 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__17(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__17); +l___auto____x40_Lean_Environment___hyg_2510____closed__18 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__18(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__18); +l___auto____x40_Lean_Environment___hyg_2510____closed__19 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__19(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__19); +l___auto____x40_Lean_Environment___hyg_2510____closed__20 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__20(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__20); +l___auto____x40_Lean_Environment___hyg_2510____closed__21 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__21(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__21); +l___auto____x40_Lean_Environment___hyg_2510____closed__22 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__22(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__22); +l___auto____x40_Lean_Environment___hyg_2510____closed__23 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__23(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__23); +l___auto____x40_Lean_Environment___hyg_2510____closed__24 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__24(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__24); +l___auto____x40_Lean_Environment___hyg_2510____closed__25 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__25(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__25); +l___auto____x40_Lean_Environment___hyg_2510____closed__26 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__26(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__26); +l___auto____x40_Lean_Environment___hyg_2510____closed__27 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__27(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__27); +l___auto____x40_Lean_Environment___hyg_2510____closed__28 = _init_l___auto____x40_Lean_Environment___hyg_2510____closed__28(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510____closed__28); +l___auto____x40_Lean_Environment___hyg_2510_ = _init_l___auto____x40_Lean_Environment___hyg_2510_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2510_); 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(); @@ -17332,16 +17334,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_2912_ = _init_l___auto____x40_Lean_Environment___hyg_2912_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2912_); +l___auto____x40_Lean_Environment___hyg_2883_ = _init_l___auto____x40_Lean_Environment___hyg_2883_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_2883_); 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_3251_ = _init_l___auto____x40_Lean_Environment___hyg_3251_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3251_); +l___auto____x40_Lean_Environment___hyg_3192_ = _init_l___auto____x40_Lean_Environment___hyg_3192_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3192_); 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(); @@ -17366,8 +17368,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_3429_ = _init_l___auto____x40_Lean_Environment___hyg_3429_(); -lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3429_); +l___auto____x40_Lean_Environment___hyg_3340_ = _init_l___auto____x40_Lean_Environment___hyg_3340_(); +lean_mark_persistent(l___auto____x40_Lean_Environment___hyg_3340_); 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(); @@ -17440,21 +17442,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_6417____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____lambda__2___closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__1); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__2(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__2); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__3(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__3); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__4); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__5); -l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6417____closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6417_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____lambda__2___closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__1); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__2); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__3 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__3(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__3); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__4); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__5); +l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__6 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__6(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6298____closed__6); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6298_(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); @@ -17466,9 +17468,9 @@ l_Lean_Kernel_instInhabitedDiagnostics___closed__1 = _init_l_Lean_Kernel_instInh lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics___closed__1); l_Lean_Kernel_instInhabitedDiagnostics = _init_l_Lean_Kernel_instInhabitedDiagnostics(); lean_mark_persistent(l_Lean_Kernel_instInhabitedDiagnostics); -l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1(); -lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6541____closed__1); -if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6541_(lean_io_mk_world()); +l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1 = _init_l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1(); +lean_mark_persistent(l_Lean_initFn____x40_Lean_Environment___hyg_6422____closed__1); +if (builtin) {res = l_Lean_initFn____x40_Lean_Environment___hyg_6422_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_diagExt = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_diagExt); diff --git a/stage0/stdlib/Lean/KeyedDeclsAttribute.c b/stage0/stdlib/Lean/KeyedDeclsAttribute.c index a21b2ca624f9..681648366964 100644 --- a/stage0/stdlib/Lean/KeyedDeclsAttribute.c +++ b/stage0/stdlib/Lean/KeyedDeclsAttribute.c @@ -102,6 +102,7 @@ LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_Def_evalKey___default___rarg LEAN_EXPORT lean_object* l_Lean_Elab_addConstInfo___at_Lean_KeyedDeclsAttribute_Def_evalKey___default___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__20___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__6(lean_object*, lean_object*); +lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_KeyedDeclsAttribute_mkStateOfTable___spec__5___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_KeyedDeclsAttribute_getEntries___spec__6(lean_object*); extern lean_object* l_Lean_LocalContext_empty; @@ -134,6 +135,7 @@ LEAN_EXPORT lean_object* l_Lean_SMap_fold___at_Lean_KeyedDeclsAttribute_mkStateO LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_ExtensionState_erase___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_KeyedDeclsAttribute_instInhabitedOLeanEntry___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_KeyedDeclsAttribute_init___spec__4(lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__30___rarg(lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__22(lean_object*); @@ -186,7 +188,6 @@ lean_object* lean_decl_get_sorry_dep(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_SMap_insert___at___private_Lean_KeyedDeclsAttribute_0__Lean_KeyedDeclsAttribute_Table_insert___spec__22___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_KeyedDeclsAttribute_getEntries___spec__5(lean_object*); lean_object* l_Lean_Attribute_Builtin_getIdent(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -285,7 +286,6 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_KeyedDeclsAttribut static lean_object* l___auto____x40_Lean_KeyedDeclsAttribute___hyg_737____closed__19; lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterTR_loop___at_Lean_KeyedDeclsAttribute_getEntries___spec__13(lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___closed__4; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_KeyedDeclsAttribute_init___spec__5___rarg(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_evalConstCheck___at_Lean_KeyedDeclsAttribute_init___spec__2(lean_object*); @@ -4866,7 +4866,7 @@ lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); lean_dec(x_1); -x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_4, x_3); return x_5; } } @@ -4885,7 +4885,7 @@ lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); lean_dec(x_1); -x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_4, x_3); return x_5; } } @@ -4936,7 +4936,7 @@ lean_object* x_4; lean_object* x_5; x_4 = lean_ctor_get(x_1, 0); lean_inc(x_4); lean_dec(x_1); -x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(x_2, x_4, x_3); +x_5 = l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(x_2, x_4, x_3); return x_5; } } @@ -6735,7 +6735,7 @@ LEAN_EXPORT lean_object* l_Lean_KeyedDeclsAttribute_init___rarg___lambda__10(lea { uint8_t x_12; uint8_t x_13; x_12 = 0; -x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_8, x_12); +x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_8, 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; uint8_t x_20; diff --git a/stage0/stdlib/Lean/LabelAttribute.c b/stage0/stdlib/Lean/LabelAttribute.c index b838acec5e0e..abe65aa4c3f2 100644 --- a/stage0/stdlib/Lean/LabelAttribute.c +++ b/stage0/stdlib/Lean/LabelAttribute.c @@ -117,6 +117,7 @@ lean_object* l_Lean_Syntax_node3(lean_object*, lean_object*, lean_object*, lean_ lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__6; +LEAN_EXPORT lean_object* l___auto____x40_Lean_LabelAttribute___hyg_160_; static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__12; static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__27; lean_object* l_Lean_addMacroScope(lean_object*, lean_object*, lean_object*); @@ -137,6 +138,7 @@ static lean_object* l___auto____x40_Lean_LabelAttribute___hyg_52____closed__8; lean_object* l_Lean_TSyntax_getDocString(lean_object*); static lean_object* l_Lean_mkLabelExt___closed__1; LEAN_EXPORT lean_object* l_Lean_mkLabelAttr___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_LabelAttribute___hyg_88_; static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__55; static lean_object* l___auto____x40_Lean_LabelAttribute___hyg_52____closed__28; size_t lean_hashmap_mk_idx(lean_object*, uint64_t); @@ -199,7 +201,6 @@ lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l_Lean_Parser_Command_registerLabelAttr___closed__12; static lean_object* l_Lean_Parser_Command_registerLabelAttr___closed__17; lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_LabelAttribute___hyg_220_; static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__8; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____root____Lean__Parser__Command__registerLabelAttr__1___closed__25; @@ -207,7 +208,6 @@ static lean_object* l_Lean___aux__Lean__LabelAttribute______macroRules__Lean____ lean_object* lean_string_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_LabelAttribute___hyg_118_; static lean_object* l___auto____x40_Lean_LabelAttribute___hyg_52____closed__2; static lean_object* l___auto____x40_Lean_LabelAttribute___hyg_52____closed__24; static lean_object* l___auto____x40_Lean_LabelAttribute___hyg_52____closed__20; @@ -636,7 +636,7 @@ x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_2); return x_7; } } -static lean_object* _init_l___auto____x40_Lean_LabelAttribute___hyg_118_() { +static lean_object* _init_l___auto____x40_Lean_LabelAttribute___hyg_88_() { _start: { lean_object* x_1; @@ -1008,7 +1008,7 @@ lean_dec(x_1); return x_6; } } -static lean_object* _init_l___auto____x40_Lean_LabelAttribute___hyg_220_() { +static lean_object* _init_l___auto____x40_Lean_LabelAttribute___hyg_160_() { _start: { lean_object* x_1; @@ -2889,8 +2889,8 @@ l_Lean_mkLabelExt___closed__1 = _init_l_Lean_mkLabelExt___closed__1(); lean_mark_persistent(l_Lean_mkLabelExt___closed__1); l_Lean_mkLabelExt___closed__2 = _init_l_Lean_mkLabelExt___closed__2(); lean_mark_persistent(l_Lean_mkLabelExt___closed__2); -l___auto____x40_Lean_LabelAttribute___hyg_118_ = _init_l___auto____x40_Lean_LabelAttribute___hyg_118_(); -lean_mark_persistent(l___auto____x40_Lean_LabelAttribute___hyg_118_); +l___auto____x40_Lean_LabelAttribute___hyg_88_ = _init_l___auto____x40_Lean_LabelAttribute___hyg_88_(); +lean_mark_persistent(l___auto____x40_Lean_LabelAttribute___hyg_88_); l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__1 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__1(); lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__1); l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__2 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__2(); @@ -2899,8 +2899,8 @@ l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__3 = _ini lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__3); l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__4 = _init_l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__4(); lean_mark_persistent(l_Lean_ScopedEnvExtension_add___at_Lean_mkLabelAttr___spec__1___closed__4); -l___auto____x40_Lean_LabelAttribute___hyg_220_ = _init_l___auto____x40_Lean_LabelAttribute___hyg_220_(); -lean_mark_persistent(l___auto____x40_Lean_LabelAttribute___hyg_220_); +l___auto____x40_Lean_LabelAttribute___hyg_160_ = _init_l___auto____x40_Lean_LabelAttribute___hyg_160_(); +lean_mark_persistent(l___auto____x40_Lean_LabelAttribute___hyg_160_); l_Lean_registerLabelAttr___closed__1 = _init_l_Lean_registerLabelAttr___closed__1(); lean_mark_persistent(l_Lean_registerLabelAttr___closed__1); l_Lean_Parser_Command_registerLabelAttr___closed__1 = _init_l_Lean_Parser_Command_registerLabelAttr___closed__1(); diff --git a/stage0/stdlib/Lean/Language/Basic.c b/stage0/stdlib/Lean/Language/Basic.c index 366f70b262b4..2ce00bec6979 100644 --- a/stage0/stdlib/Lean/Language/Basic.c +++ b/stage0/stdlib/Lean/Language/Basic.c @@ -22,11 +22,10 @@ LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTask_cancel___rarg(lean_object* lean_object* lean_mk_empty_array_with_capacity(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_printMessageEndPos; static lean_object* l_Lean_Language_instImpl____x40_Lean_Language_Basic___hyg_728____closed__1; -LEAN_EXPORT uint32_t l_Lean_Language_ModuleProcessingContext_trustLevel___default; LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTree_children(lean_object*); static lean_object* l_Lean_Language_instImpl____x40_Lean_Language_Basic___hyg_728____closed__3; LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_toTyped_x3f___rarg___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_toTyped_x3f___rarg(lean_object*, lean_object*); lean_object* lean_io_has_finished(lean_object*, lean_object*); @@ -113,7 +112,7 @@ static lean_object* l_Lean_Language_Snapshot_instInhabitedDiagnostics___closed__ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_reportMessages___spec__6(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_instToSnapshotTreeDynamicSnapshot___boxed(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_reportMessages___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_SnapshotTree_getAll___spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_SnapshotTree_getAll___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -143,7 +142,7 @@ lean_object* lean_array_uget(lean_object*, size_t); lean_object* l_Array_foldlMUnsafe_fold___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_SnapshotTask_get___rarg(lean_object*); static lean_object* l_Lean_Language_initFn____x40_Lean_Language_Basic___hyg_938____closed__3; -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_error_to_string(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_reportMessages___lambda__2(lean_object*, lean_object*); @@ -2203,14 +2202,6 @@ lean_dec(x_1); return x_8; } } -static uint32_t _init_l_Lean_Language_ModuleProcessingContext_trustLevel___default() { -_start: -{ -uint32_t x_1; -x_1 = 0; -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_Language_instMonadLiftProcessingMProcessingTIO___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -2333,36 +2324,35 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Language_diagnosticsOfHeaderError(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; -x_4 = lean_ctor_get(x_2, 1); -x_5 = lean_ctor_get(x_4, 2); -x_6 = lean_ctor_get(x_5, 0); -x_7 = lean_string_utf8_byte_size(x_6); -x_8 = l_Lean_FileMap_toPosition(x_5, x_7); -lean_dec(x_7); -x_9 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_9, 0, x_8); -x_10 = lean_alloc_ctor(3, 1, 0); -lean_ctor_set(x_10, 0, x_1); -x_11 = lean_alloc_ctor(0, 1, 0); -lean_ctor_set(x_11, 0, x_10); -x_12 = l_Lean_Language_diagnosticsOfHeaderError___closed__2; -x_13 = l_Lean_Language_diagnosticsOfHeaderError___closed__1; -x_14 = 0; -x_15 = 2; -x_16 = l_Lean_Language_initFn____x40_Lean_Language_Basic___hyg_938____closed__3; -x_17 = lean_alloc_ctor(0, 5, 2); -lean_ctor_set(x_17, 0, x_12); -lean_ctor_set(x_17, 1, x_13); -lean_ctor_set(x_17, 2, x_9); -lean_ctor_set(x_17, 3, x_16); -lean_ctor_set(x_17, 4, x_11); -lean_ctor_set_uint8(x_17, sizeof(void*)*5, x_14); -lean_ctor_set_uint8(x_17, sizeof(void*)*5 + 1, x_15); -x_18 = l_Lean_MessageLog_empty; -x_19 = l_Lean_PersistentArray_push___rarg(x_18, x_17); -x_20 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_19, x_3); -return x_20; +lean_object* x_4; lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; uint8_t x_13; uint8_t x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; +x_4 = lean_ctor_get(x_2, 2); +x_5 = lean_ctor_get(x_4, 0); +x_6 = lean_string_utf8_byte_size(x_5); +x_7 = l_Lean_FileMap_toPosition(x_4, x_6); +lean_dec(x_6); +x_8 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_8, 0, x_7); +x_9 = lean_alloc_ctor(3, 1, 0); +lean_ctor_set(x_9, 0, x_1); +x_10 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_10, 0, x_9); +x_11 = l_Lean_Language_diagnosticsOfHeaderError___closed__2; +x_12 = l_Lean_Language_diagnosticsOfHeaderError___closed__1; +x_13 = 0; +x_14 = 2; +x_15 = l_Lean_Language_initFn____x40_Lean_Language_Basic___hyg_938____closed__3; +x_16 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_16, 0, x_11); +lean_ctor_set(x_16, 1, x_12); +lean_ctor_set(x_16, 2, x_8); +lean_ctor_set(x_16, 3, x_15); +lean_ctor_set(x_16, 4, x_10); +lean_ctor_set_uint8(x_16, sizeof(void*)*5, x_13); +lean_ctor_set_uint8(x_16, sizeof(void*)*5 + 1, x_14); +x_17 = l_Lean_MessageLog_empty; +x_18 = l_Lean_PersistentArray_push___rarg(x_17, x_16); +x_19 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_18, x_3); +return x_19; } } LEAN_EXPORT lean_object* l_Lean_Language_diagnosticsOfHeaderError___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { @@ -2488,129 +2478,125 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Language_withHeaderExceptions___rarg), 4 return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg(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_Language_mkIncrementalProcessor___elambda__1___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; -x_6 = lean_st_ref_get(x_3, x_5); -if (lean_obj_tag(x_6) == 0) +lean_object* x_5; +x_5 = lean_st_ref_get(x_2, x_4); +if (lean_obj_tag(x_5) == 0) { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; -x_7 = lean_ctor_get(x_6, 0); +lean_object* x_6; lean_object* x_7; lean_object* x_8; +x_6 = lean_ctor_get(x_5, 0); +lean_inc(x_6); +x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); -x_8 = lean_ctor_get(x_6, 1); -lean_inc(x_8); -lean_dec(x_6); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_2); -lean_ctor_set(x_9, 1, x_4); -x_10 = lean_apply_3(x_1, x_7, x_9, x_8); -if (lean_obj_tag(x_10) == 0) +lean_dec(x_5); +x_8 = lean_apply_3(x_1, x_6, x_3, x_7); +if (lean_obj_tag(x_8) == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -lean_inc(x_11); -x_13 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_13, 0, x_11); -x_14 = lean_st_ref_set(x_3, x_13, x_12); -if (lean_obj_tag(x_14) == 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_inc(x_10); +lean_dec(x_8); +lean_inc(x_9); +x_11 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_11, 0, x_9); +x_12 = lean_st_ref_set(x_2, x_11, x_10); +if (lean_obj_tag(x_12) == 0) { -uint8_t x_15; -x_15 = !lean_is_exclusive(x_14); -if (x_15 == 0) +uint8_t x_13; +x_13 = !lean_is_exclusive(x_12); +if (x_13 == 0) { -lean_object* x_16; -x_16 = lean_ctor_get(x_14, 0); -lean_dec(x_16); -lean_ctor_set(x_14, 0, x_11); -return x_14; +lean_object* x_14; +x_14 = lean_ctor_get(x_12, 0); +lean_dec(x_14); +lean_ctor_set(x_12, 0, x_9); +return x_12; } else { -lean_object* x_17; lean_object* x_18; -x_17 = lean_ctor_get(x_14, 1); -lean_inc(x_17); -lean_dec(x_14); -x_18 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_18, 0, x_11); -lean_ctor_set(x_18, 1, x_17); -return x_18; +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_9); +lean_ctor_set(x_16, 1, x_15); +return x_16; } } else { -uint8_t x_19; -lean_dec(x_11); -x_19 = !lean_is_exclusive(x_14); -if (x_19 == 0) +uint8_t x_17; +lean_dec(x_9); +x_17 = !lean_is_exclusive(x_12); +if (x_17 == 0) { -return x_14; +return x_12; } else { -lean_object* x_20; lean_object* x_21; lean_object* x_22; -x_20 = lean_ctor_get(x_14, 0); -x_21 = lean_ctor_get(x_14, 1); -lean_inc(x_21); -lean_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_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; } } } else { -uint8_t x_23; -x_23 = !lean_is_exclusive(x_10); -if (x_23 == 0) +uint8_t x_21; +x_21 = !lean_is_exclusive(x_8); +if (x_21 == 0) { -return x_10; +return x_8; } else { -lean_object* x_24; lean_object* x_25; lean_object* x_26; -x_24 = lean_ctor_get(x_10, 0); -x_25 = lean_ctor_get(x_10, 1); -lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_10); -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_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; } } } else { -uint8_t x_27; -lean_dec(x_4); -lean_dec(x_2); +uint8_t x_25; +lean_dec(x_3); lean_dec(x_1); -x_27 = !lean_is_exclusive(x_6); -if (x_27 == 0) +x_25 = !lean_is_exclusive(x_5); +if (x_25 == 0) { -return x_6; +return x_5; } else { -lean_object* x_28; lean_object* x_29; lean_object* x_30; -x_28 = lean_ctor_get(x_6, 0); -x_29 = lean_ctor_get(x_6, 1); -lean_inc(x_29); -lean_inc(x_28); -lean_dec(x_6); -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_26; lean_object* x_27; lean_object* x_28; +x_26 = lean_ctor_get(x_5, 0); +x_27 = lean_ctor_get(x_5, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_5); +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; } } } @@ -2619,71 +2605,68 @@ LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1(lea _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 5, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 4, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object* x_1, lean_object* x_2) { _start: { -lean_object* x_4; lean_object* x_5; -x_4 = lean_box(0); -x_5 = lean_st_mk_ref(x_4, x_3); -if (lean_obj_tag(x_5) == 0) +lean_object* x_3; lean_object* x_4; +x_3 = lean_box(0); +x_4 = lean_st_mk_ref(x_3, x_2); +if (lean_obj_tag(x_4) == 0) { -uint8_t x_6; -x_6 = !lean_is_exclusive(x_5); -if (x_6 == 0) +uint8_t x_5; +x_5 = !lean_is_exclusive(x_4); +if (x_5 == 0) { -lean_object* x_7; lean_object* x_8; -x_7 = lean_ctor_get(x_5, 0); -x_8 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 5, 3); -lean_closure_set(x_8, 0, x_1); -lean_closure_set(x_8, 1, x_2); -lean_closure_set(x_8, 2, x_7); -lean_ctor_set(x_5, 0, x_8); -return x_5; +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_4, 0); +x_7 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 4, 2); +lean_closure_set(x_7, 0, x_1); +lean_closure_set(x_7, 1, x_6); +lean_ctor_set(x_4, 0, x_7); +return x_4; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 0); -x_10 = lean_ctor_get(x_5, 1); -lean_inc(x_10); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_4, 0); +x_9 = lean_ctor_get(x_4, 1); lean_inc(x_9); -lean_dec(x_5); -x_11 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 5, 3); -lean_closure_set(x_11, 0, x_1); -lean_closure_set(x_11, 1, x_2); -lean_closure_set(x_11, 2, x_9); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_11); -lean_ctor_set(x_12, 1, x_10); -return x_12; +lean_inc(x_8); +lean_dec(x_4); +x_10 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed), 4, 2); +lean_closure_set(x_10, 0, x_1); +lean_closure_set(x_10, 1, x_8); +x_11 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_11, 0, x_10); +lean_ctor_set(x_11, 1, x_9); +return x_11; } } else { -uint8_t x_13; -lean_dec(x_2); +uint8_t x_12; lean_dec(x_1); -x_13 = !lean_is_exclusive(x_5); -if (x_13 == 0) +x_12 = !lean_is_exclusive(x_4); +if (x_12 == 0) { -return x_5; +return x_4; } else { -lean_object* x_14; lean_object* x_15; lean_object* x_16; -x_14 = lean_ctor_get(x_5, 0); -x_15 = lean_ctor_get(x_5, 1); -lean_inc(x_15); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +x_13 = lean_ctor_get(x_4, 0); +x_14 = lean_ctor_get(x_4, 1); lean_inc(x_14); -lean_dec(x_5); -x_16 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_16, 0, x_14); -lean_ctor_set(x_16, 1, x_15); -return x_16; +lean_inc(x_13); +lean_dec(x_4); +x_15 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_15, 0, x_13); +lean_ctor_set(x_15, 1, x_14); +return x_15; } } } @@ -2692,17 +2675,17 @@ LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor(lean_object* x_1 _start: { lean_object* x_2; -x_2 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___rarg), 3, 0); +x_2 = lean_alloc_closure((void*)(l_Lean_Language_mkIncrementalProcessor___rarg), 2, 0); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Language_mkIncrementalProcessor___elambda__1___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_Lean_Language_mkIncrementalProcessor___elambda__1___rarg___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_6; -x_6 = l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_3); -return x_6; +lean_object* x_5; +x_5 = l_Lean_Language_mkIncrementalProcessor___elambda__1___rarg(x_1, x_2, x_3, x_4); +lean_dec(x_2); +return x_5; } } lean_object* initialize_Init_System_Promise(uint8_t builtin, lean_object*); @@ -2783,7 +2766,6 @@ l_Lean_Language_reportMessages___lambda__2___closed__1 = _init_l_Lean_Language_r lean_mark_persistent(l_Lean_Language_reportMessages___lambda__2___closed__1); l_Lean_Language_reportMessages___closed__1 = _init_l_Lean_Language_reportMessages___closed__1(); lean_mark_persistent(l_Lean_Language_reportMessages___closed__1); -l_Lean_Language_ModuleProcessingContext_trustLevel___default = _init_l_Lean_Language_ModuleProcessingContext_trustLevel___default(); l_Lean_Language_diagnosticsOfHeaderError___closed__1 = _init_l_Lean_Language_diagnosticsOfHeaderError___closed__1(); lean_mark_persistent(l_Lean_Language_diagnosticsOfHeaderError___closed__1); l_Lean_Language_diagnosticsOfHeaderError___closed__2 = _init_l_Lean_Language_diagnosticsOfHeaderError___closed__2(); diff --git a/stage0/stdlib/Lean/Language/Lean.c b/stage0/stdlib/Lean/Language/Lean.c index f2662a36856d..7b86aaa48ca3 100644 --- a/stage0/stdlib/Lean/Language/Lean.c +++ b/stage0/stdlib/Lean/Language/Lean.c @@ -13,27 +13,27 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_Language_Lean_process(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; LEAN_EXPORT lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_71_(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___closed__5; static lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__5___closed__1; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT uint32_t l_Lean_Language_Lean_SetupImportsResult_trustLevel___default; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___closed__7; static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___closed__3; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_SnapshotTask_cancel___rarg(lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); size_t lean_usize_shift_right(size_t, size_t); -lean_object* l_List_forIn_loop___at_Lean_KVMap_mergeBy___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Language_Lean_HeaderProcessedSnapshot_isFatal___default(lean_object*); lean_object* l_IO_mkRef___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_CommandParsedSnapshot_cancel___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandFinishedSnapshot___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeHeaderProcessedSnapshot(lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10; static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__4; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8; lean_object* lean_io_check_canceled(lean_object*); lean_object* l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_DynamicSnapshot_ofTyped___at_Lean_Language_Lean_process_doElab___spec__5(lean_object*, lean_object*); @@ -49,11 +49,12 @@ static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_ LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Elab_Command_elabCommandTopLevel(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_36____closed__4; LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f(lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___closed__8; @@ -67,13 +68,14 @@ lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); lean_object* l_Lean_Language_SnapshotTask_ofIO___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_process_parseCmd___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; LEAN_EXPORT lean_object* l_Lean_Language_Lean_CommandParsedSnapshot_cancel(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___boxed(lean_object*); static lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___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_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_CommandParsedSnapshot_next_x3f(lean_object*); lean_object* l_List_get___rarg(lean_object*, lean_object*); @@ -81,6 +83,7 @@ lean_object* l_Lean_Elab_processHeader(lean_object*, lean_object*, lean_object*, lean_object* lean_io_promise_result(lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1___closed__1; lean_object* l_Lean_instInhabitedPersistentArrayNode(lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; extern lean_object* l_Lean_maxRecDepth; lean_object* l_List_head_x21___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__2; @@ -94,18 +97,16 @@ uint8_t lean_string_validate_utf8(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go___lambda__1___boxed(lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; LEAN_EXPORT lean_object* l_Lean_Language_Lean_CommandParsedSnapshot_data(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot; lean_object* lean_io_promise_resolve(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11; extern lean_object* l_Lean_Elab_instInhabitedInfoTree; size_t lean_usize_of_nat(lean_object*); uint8_t l_String_isEmpty(lean_object*); uint8_t l_Lean_Parser_isTerminalCommand(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_waitForFinalEnv_x3f_goCmd(lean_object*); static lean_object* l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__1; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6; LEAN_EXPORT uint8_t l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10___lambda__1(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_36____closed__8; @@ -115,10 +116,11 @@ LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__4___box static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__8; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__6; extern lean_object* l_ByteArray_empty; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_toPArray_x27___rarg(lean_object*); lean_object* l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_71____closed__4; +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_36____closed__2; LEAN_EXPORT lean_object* l_Lean_Language_Lean_CommandParsedSnapshot_next_x3f___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeHeaderParsedSnapshot___lambda__1(lean_object*, lean_object*); @@ -137,7 +139,6 @@ LEAN_EXPORT lean_object* l_Lean_Language_Lean_HeaderParsedSnapshot_processedResu lean_object* l_String_firstDiffPos(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9; static lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go___closed__3; LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go___closed__2; @@ -146,7 +147,6 @@ static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_ lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* lean_array_to_list(lean_object*, lean_object*); uint8_t l_Lean_MessageData_hasTag(lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_36____closed__5; lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); lean_object* l_Array_toPArray_x27___rarg(lean_object*); @@ -181,9 +181,7 @@ static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_do LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__2(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_instToSnapshotTreeHeaderProcessedSnapshot___closed__1; lean_object* lean_string_from_utf8(lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_Language_Lean_process_processHeader___lambda__4___closed__4; +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go___lambda__1(lean_object*); lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandFinishedSnapshot___boxed(lean_object*); @@ -195,21 +193,19 @@ uint8_t l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(lea LEAN_EXPORT lean_object* l_Lean_PersistentArray_foldlM___at_Lean_Language_Lean_process_doElab___spec__6(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_doElab___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandFinishedSnapshot(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__2; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_71____closed__5; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Language_Lean_0__Lean_Language_Lean_pushOpt___rarg(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Language_Lean_HeaderParsedSnapshot_isFatal___default(lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12; static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__7; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_71____closed__6; -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_instMonadLiftLeanProcessingMLeanProcessingTIO___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__7___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_isBeforeEditPos(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(lean_object*, uint8_t, lean_object*); @@ -224,6 +220,7 @@ LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Language_Lean_process_processH lean_object* l_Lean_Syntax_getRange_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; LEAN_EXPORT lean_object* l_Lean_Language_Lean_isBeforeEditPos___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_71____closed__2; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); @@ -232,6 +229,7 @@ LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__5___ static lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go___closed__4; static lean_object* l_Lean_Language_Lean_initFn____x40_Lean_Language_Lean___hyg_36____closed__3; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__10(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__9(lean_object*, size_t, size_t, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -266,7 +264,6 @@ uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Language_Lean_process_doElab___lambda__4___closed__3; lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_EStateM_bind___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Language_Lean_process_doElab___spec__10(lean_object*, size_t, size_t, lean_object*); lean_object* l_Lean_Parser_parseCommand(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader___lambda__7(lean_object*, lean_object*, lean_object*, lean_object*); @@ -1422,6 +1419,14 @@ x_2 = lean_alloc_closure((void*)(l___private_Lean_Language_Lean_0__Lean_Language return x_2; } } +static uint32_t _init_l_Lean_Language_Lean_SetupImportsResult_trustLevel___default() { +_start: +{ +uint32_t x_1; +x_1 = 0; +return x_1; +} +} LEAN_EXPORT lean_object* l_IO_withStdout___at_Lean_Language_Lean_process_doElab___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { @@ -3181,7 +3186,7 @@ lean_dec(x_10); x_13 = lean_st_mk_ref(x_11, x_12); if (lean_obj_tag(x_13) == 0) { -lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; uint8_t x_30; lean_object* x_31; x_14 = lean_ctor_get(x_13, 0); lean_inc(x_14); x_15 = lean_ctor_get(x_13, 1); @@ -3192,384 +3197,381 @@ lean_inc(x_16); lean_dec(x_2); x_17 = lean_ctor_get(x_16, 1); lean_inc(x_17); -lean_dec(x_16); -x_18 = lean_ctor_get(x_17, 1); +x_18 = lean_ctor_get(x_16, 2); lean_inc(x_18); -x_19 = lean_ctor_get(x_17, 2); -lean_inc(x_19); -lean_dec(x_17); -x_20 = lean_box(0); +lean_dec(x_16); +x_19 = lean_box(0); lean_inc(x_14); -x_21 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_21, 0, x_14); +x_20 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_20, 0, x_14); lean_inc(x_3); -x_22 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_22, 0, x_3); -x_23 = lean_unsigned_to_nat(0u); -x_24 = l_Lean_firstFrontendMacroScope; -x_25 = lean_box(0); +x_21 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_21, 0, x_3); +x_22 = lean_unsigned_to_nat(0u); +x_23 = l_Lean_firstFrontendMacroScope; +x_24 = lean_box(0); lean_inc(x_4); -lean_inc(x_19); lean_inc(x_18); -x_26 = lean_alloc_ctor(0, 9, 0); -lean_ctor_set(x_26, 0, x_18); -lean_ctor_set(x_26, 1, x_19); -lean_ctor_set(x_26, 2, x_23); -lean_ctor_set(x_26, 3, x_4); -lean_ctor_set(x_26, 4, x_20); -lean_ctor_set(x_26, 5, x_24); -lean_ctor_set(x_26, 6, x_25); -lean_ctor_set(x_26, 7, x_21); -lean_ctor_set(x_26, 8, x_22); +lean_inc(x_17); +x_25 = lean_alloc_ctor(0, 9, 0); +lean_ctor_set(x_25, 0, x_17); +lean_ctor_set(x_25, 1, x_18); +lean_ctor_set(x_25, 2, x_22); +lean_ctor_set(x_25, 3, x_4); +lean_ctor_set(x_25, 4, x_19); +lean_ctor_set(x_25, 5, x_23); +lean_ctor_set(x_25, 6, x_24); +lean_ctor_set(x_25, 7, x_20); +lean_ctor_set(x_25, 8, x_21); lean_inc(x_5); -x_27 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__1), 4, 1); -lean_closure_set(x_27, 0, x_5); +x_26 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_doElab___lambda__1), 4, 1); +lean_closure_set(x_26, 0, x_5); lean_inc(x_8); -x_28 = lean_alloc_closure((void*)(l_Lean_Elab_Command_catchExceptions), 4, 3); -lean_closure_set(x_28, 0, x_27); -lean_closure_set(x_28, 1, x_26); -lean_closure_set(x_28, 2, x_8); -x_29 = lean_ctor_get(x_6, 1); -lean_inc(x_29); +x_27 = lean_alloc_closure((void*)(l_Lean_Elab_Command_catchExceptions), 4, 3); +lean_closure_set(x_27, 0, x_26); +lean_closure_set(x_27, 1, x_25); +lean_closure_set(x_27, 2, x_8); +x_28 = lean_ctor_get(x_6, 1); +lean_inc(x_28); lean_dec(x_6); -x_30 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__1; -x_31 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_29, x_30); -lean_dec(x_29); -x_32 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_28, x_31, x_15); -if (lean_obj_tag(x_32) == 0) +x_29 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__1; +x_30 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_28, x_29); +lean_dec(x_28); +x_31 = l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1(x_27, x_30, x_15); +if (lean_obj_tag(x_31) == 0) { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; -x_33 = lean_ctor_get(x_32, 0); +lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; +x_32 = lean_ctor_get(x_31, 0); +lean_inc(x_32); +x_33 = lean_ctor_get(x_31, 1); lean_inc(x_33); -x_34 = lean_ctor_get(x_32, 1); +lean_dec(x_31); +x_34 = lean_ctor_get(x_32, 0); lean_inc(x_34); lean_dec(x_32); -x_35 = lean_ctor_get(x_33, 0); -lean_inc(x_35); -lean_dec(x_33); -x_36 = lean_st_ref_get(x_14, x_34); +x_35 = lean_st_ref_get(x_14, x_33); lean_dec(x_14); -if (lean_obj_tag(x_36) == 0) +if (lean_obj_tag(x_35) == 0) { -lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_37 = lean_ctor_get(x_36, 0); +lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); lean_inc(x_37); +lean_dec(x_35); x_38 = lean_ctor_get(x_36, 1); lean_inc(x_38); lean_dec(x_36); -x_39 = lean_ctor_get(x_37, 1); -lean_inc(x_39); -lean_dec(x_37); -x_40 = lean_st_ref_take(x_1, x_38); -if (lean_obj_tag(x_40) == 0) +x_39 = lean_st_ref_take(x_1, x_37); +if (lean_obj_tag(x_39) == 0) { -lean_object* x_41; lean_object* x_42; lean_object* x_43; lean_object* x_44; -x_41 = lean_ctor_get(x_40, 1); -lean_inc(x_41); -lean_dec(x_40); -x_42 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__4; -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_st_ref_set(x_1, x_43, x_41); +lean_object* x_40; lean_object* x_41; lean_object* x_42; lean_object* x_43; +x_40 = lean_ctor_get(x_39, 1); +lean_inc(x_40); +lean_dec(x_39); +x_41 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__4; +x_42 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_42, 0, x_38); +lean_ctor_set(x_42, 1, x_41); +x_43 = lean_st_ref_set(x_1, x_42, x_40); lean_dec(x_1); -if (lean_obj_tag(x_44) == 0) +if (lean_obj_tag(x_43) == 0) { -lean_object* x_45; lean_object* x_46; -x_45 = lean_ctor_get(x_44, 1); -lean_inc(x_45); -lean_dec(x_44); -x_46 = lean_st_ref_get(x_8, x_45); +lean_object* x_44; lean_object* x_45; +x_44 = lean_ctor_get(x_43, 1); +lean_inc(x_44); +lean_dec(x_43); +x_45 = lean_st_ref_get(x_8, x_44); lean_dec(x_8); -if (lean_obj_tag(x_46) == 0) +if (lean_obj_tag(x_45) == 0) { -lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; uint8_t x_52; -x_47 = lean_ctor_get(x_46, 0); +lean_object* x_46; lean_object* x_47; lean_object* x_48; lean_object* x_49; lean_object* x_50; uint8_t x_51; +x_46 = lean_ctor_get(x_45, 0); +lean_inc(x_46); +x_47 = lean_ctor_get(x_45, 1); lean_inc(x_47); +lean_dec(x_45); x_48 = lean_ctor_get(x_46, 1); lean_inc(x_48); -lean_dec(x_46); -x_49 = lean_ctor_get(x_47, 1); +x_49 = lean_ctor_get(x_46, 2); lean_inc(x_49); -x_50 = lean_ctor_get(x_47, 2); -lean_inc(x_50); -x_51 = l_List_lengthTRAux___rarg(x_50, x_23); -x_52 = lean_nat_dec_lt(x_23, x_51); -lean_dec(x_51); -if (x_52 == 0) -{ -lean_object* x_53; lean_object* x_54; lean_object* x_55; lean_object* x_56; uint8_t x_57; +x_50 = l_List_lengthTRAux___rarg(x_49, x_22); +x_51 = lean_nat_dec_lt(x_22, x_50); lean_dec(x_50); -x_53 = l_Lean_Elab_Command_instInhabitedScope; -x_54 = l___private_Init_GetElem_0__outOfBounds___rarg(x_53); -x_55 = lean_ctor_get(x_54, 1); -lean_inc(x_55); +if (x_51 == 0) +{ +lean_object* x_52; lean_object* x_53; lean_object* x_54; lean_object* x_55; uint8_t x_56; +lean_dec(x_49); +x_52 = l_Lean_Elab_Command_instInhabitedScope; +x_53 = l___private_Init_GetElem_0__outOfBounds___rarg(x_52); +x_54 = lean_ctor_get(x_53, 1); +lean_inc(x_54); +lean_dec(x_53); +x_55 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__8; +x_56 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_54, x_55); lean_dec(x_54); -x_56 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__8; -x_57 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_55, x_56); -lean_dec(x_55); -if (x_57 == 0) +if (x_56 == 0) { if (x_7 == 0) { -lean_object* x_58; lean_object* x_59; +lean_object* x_57; lean_object* x_58; lean_dec(x_5); -x_58 = lean_box(0); -x_59 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_58, x_48); +x_57 = lean_box(0); +x_58 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_57, x_47); lean_dec(x_3); -return x_59; +return x_58; } else { -lean_object* x_60; uint8_t x_61; -x_60 = l_Lean_Syntax_hasMissing(x_5); -x_61 = lean_unbox(x_60); -lean_dec(x_60); -if (x_61 == 0) +lean_object* x_59; uint8_t x_60; +x_59 = l_Lean_Syntax_hasMissing(x_5); +x_60 = lean_unbox(x_59); +lean_dec(x_59); +if (x_60 == 0) { -lean_object* x_62; lean_object* x_63; -x_62 = lean_box(0); -x_63 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_62, x_48); +lean_object* x_61; lean_object* x_62; +x_61 = lean_box(0); +x_62 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_61, x_47); lean_dec(x_3); -return x_63; +return x_62; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_64 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__7; -x_65 = l_Lean_PersistentArray_foldlM___at_Lean_Language_Lean_process_doElab___spec__6(x_49, x_64, x_23); -x_66 = lean_box(0); -x_67 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_65, x_66, x_48); +lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_63 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__7; +x_64 = l_Lean_PersistentArray_foldlM___at_Lean_Language_Lean_process_doElab___spec__6(x_48, x_63, x_22); +x_65 = lean_box(0); +x_66 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_64, x_65, x_47); lean_dec(x_3); -return x_67; +return x_66; } } } else { -lean_object* x_68; lean_object* x_69; +lean_object* x_67; lean_object* x_68; lean_dec(x_5); -x_68 = lean_box(0); -x_69 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_68, x_48); +x_67 = lean_box(0); +x_68 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_67, x_47); lean_dec(x_3); -return x_69; +return x_68; } } else { -lean_object* x_70; lean_object* x_71; lean_object* x_72; uint8_t x_73; -x_70 = l_List_get___rarg(x_50, x_23); -lean_dec(x_50); -x_71 = lean_ctor_get(x_70, 1); -lean_inc(x_71); +lean_object* x_69; lean_object* x_70; lean_object* x_71; uint8_t x_72; +x_69 = l_List_get___rarg(x_49, x_22); +lean_dec(x_49); +x_70 = lean_ctor_get(x_69, 1); +lean_inc(x_70); +lean_dec(x_69); +x_71 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__8; +x_72 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_70, x_71); lean_dec(x_70); -x_72 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__8; -x_73 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_71, x_72); -lean_dec(x_71); -if (x_73 == 0) +if (x_72 == 0) { if (x_7 == 0) { -lean_object* x_74; lean_object* x_75; +lean_object* x_73; lean_object* x_74; lean_dec(x_5); -x_74 = lean_box(0); -x_75 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_74, x_48); +x_73 = lean_box(0); +x_74 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_73, x_47); lean_dec(x_3); -return x_75; +return x_74; } else { -lean_object* x_76; uint8_t x_77; -x_76 = l_Lean_Syntax_hasMissing(x_5); -x_77 = lean_unbox(x_76); -lean_dec(x_76); -if (x_77 == 0) +lean_object* x_75; uint8_t x_76; +x_75 = l_Lean_Syntax_hasMissing(x_5); +x_76 = lean_unbox(x_75); +lean_dec(x_75); +if (x_76 == 0) { -lean_object* x_78; lean_object* x_79; -x_78 = lean_box(0); -x_79 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_78, x_48); +lean_object* x_77; lean_object* x_78; +x_77 = lean_box(0); +x_78 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_77, x_47); lean_dec(x_3); -return x_79; +return x_78; } else { -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -x_80 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__7; -x_81 = l_Lean_PersistentArray_foldlM___at_Lean_Language_Lean_process_doElab___spec__6(x_49, x_80, x_23); -x_82 = lean_box(0); -x_83 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_81, x_82, x_48); +lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; +x_79 = l_Lean_Language_Lean_process_doElab___lambda__4___closed__7; +x_80 = l_Lean_PersistentArray_foldlM___at_Lean_Language_Lean_process_doElab___spec__6(x_48, x_79, x_22); +x_81 = lean_box(0); +x_82 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_80, x_81, x_47); lean_dec(x_3); -return x_83; +return x_82; } } } else { -lean_object* x_84; lean_object* x_85; +lean_object* x_83; lean_object* x_84; lean_dec(x_5); -x_84 = lean_box(0); -x_85 = l_Lean_Language_Lean_process_doElab___lambda__3(x_47, x_3, x_35, x_19, x_4, x_18, x_49, x_84, x_48); +x_83 = lean_box(0); +x_84 = l_Lean_Language_Lean_process_doElab___lambda__3(x_46, x_3, x_34, x_18, x_4, x_17, x_48, x_83, x_47); lean_dec(x_3); -return x_85; +return x_84; } } } else { -uint8_t x_86; -lean_dec(x_35); -lean_dec(x_19); +uint8_t x_85; +lean_dec(x_34); lean_dec(x_18); +lean_dec(x_17); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_86 = !lean_is_exclusive(x_46); -if (x_86 == 0) +x_85 = !lean_is_exclusive(x_45); +if (x_85 == 0) { -return x_46; +return x_45; } else { -lean_object* x_87; lean_object* x_88; lean_object* x_89; -x_87 = lean_ctor_get(x_46, 0); -x_88 = lean_ctor_get(x_46, 1); -lean_inc(x_88); +lean_object* x_86; lean_object* x_87; lean_object* x_88; +x_86 = lean_ctor_get(x_45, 0); +x_87 = lean_ctor_get(x_45, 1); lean_inc(x_87); -lean_dec(x_46); -x_89 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_89, 0, x_87); -lean_ctor_set(x_89, 1, x_88); -return x_89; +lean_inc(x_86); +lean_dec(x_45); +x_88 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_88, 0, x_86); +lean_ctor_set(x_88, 1, x_87); +return x_88; } } } else { -uint8_t x_90; -lean_dec(x_35); -lean_dec(x_19); +uint8_t x_89; +lean_dec(x_34); lean_dec(x_18); +lean_dec(x_17); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); -x_90 = !lean_is_exclusive(x_44); -if (x_90 == 0) +x_89 = !lean_is_exclusive(x_43); +if (x_89 == 0) { -return x_44; +return x_43; } else { -lean_object* x_91; lean_object* x_92; lean_object* x_93; -x_91 = lean_ctor_get(x_44, 0); -x_92 = lean_ctor_get(x_44, 1); -lean_inc(x_92); +lean_object* x_90; lean_object* x_91; lean_object* x_92; +x_90 = lean_ctor_get(x_43, 0); +x_91 = lean_ctor_get(x_43, 1); lean_inc(x_91); -lean_dec(x_44); -x_93 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_93, 0, x_91); -lean_ctor_set(x_93, 1, x_92); -return x_93; +lean_inc(x_90); +lean_dec(x_43); +x_92 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_92, 0, x_90); +lean_ctor_set(x_92, 1, x_91); +return x_92; } } } else { -uint8_t x_94; -lean_dec(x_39); -lean_dec(x_35); -lean_dec(x_19); +uint8_t x_93; +lean_dec(x_38); +lean_dec(x_34); lean_dec(x_18); +lean_dec(x_17); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_94 = !lean_is_exclusive(x_40); -if (x_94 == 0) +x_93 = !lean_is_exclusive(x_39); +if (x_93 == 0) { -return x_40; +return x_39; } else { -lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_95 = lean_ctor_get(x_40, 0); -x_96 = lean_ctor_get(x_40, 1); -lean_inc(x_96); +lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_94 = lean_ctor_get(x_39, 0); +x_95 = lean_ctor_get(x_39, 1); lean_inc(x_95); -lean_dec(x_40); -x_97 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_97, 0, x_95); -lean_ctor_set(x_97, 1, x_96); -return x_97; +lean_inc(x_94); +lean_dec(x_39); +x_96 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_96, 0, x_94); +lean_ctor_set(x_96, 1, x_95); +return x_96; } } } else { -uint8_t x_98; -lean_dec(x_35); -lean_dec(x_19); +uint8_t x_97; +lean_dec(x_34); lean_dec(x_18); +lean_dec(x_17); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_98 = !lean_is_exclusive(x_36); -if (x_98 == 0) +x_97 = !lean_is_exclusive(x_35); +if (x_97 == 0) { -return x_36; +return x_35; } else { -lean_object* x_99; lean_object* x_100; lean_object* x_101; -x_99 = lean_ctor_get(x_36, 0); -x_100 = lean_ctor_get(x_36, 1); -lean_inc(x_100); +lean_object* x_98; lean_object* x_99; lean_object* x_100; +x_98 = lean_ctor_get(x_35, 0); +x_99 = lean_ctor_get(x_35, 1); lean_inc(x_99); -lean_dec(x_36); -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; +lean_inc(x_98); +lean_dec(x_35); +x_100 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_100, 0, x_98); +lean_ctor_set(x_100, 1, x_99); +return x_100; } } } else { -uint8_t x_102; -lean_dec(x_19); +uint8_t x_101; lean_dec(x_18); +lean_dec(x_17); lean_dec(x_14); lean_dec(x_8); lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_1); -x_102 = !lean_is_exclusive(x_32); -if (x_102 == 0) +x_101 = !lean_is_exclusive(x_31); +if (x_101 == 0) { -return x_32; +return x_31; } else { -lean_object* x_103; lean_object* x_104; lean_object* x_105; -x_103 = lean_ctor_get(x_32, 0); -x_104 = lean_ctor_get(x_32, 1); -lean_inc(x_104); +lean_object* x_102; lean_object* x_103; lean_object* x_104; +x_102 = lean_ctor_get(x_31, 0); +x_103 = lean_ctor_get(x_31, 1); lean_inc(x_103); -lean_dec(x_32); -x_105 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_105, 0, x_103); -lean_ctor_set(x_105, 1, x_104); -return x_105; +lean_inc(x_102); +lean_dec(x_31); +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_106; +uint8_t x_105; lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -3577,29 +3579,29 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_106 = !lean_is_exclusive(x_13); -if (x_106 == 0) +x_105 = !lean_is_exclusive(x_13); +if (x_105 == 0) { return x_13; } else { -lean_object* x_107; lean_object* x_108; lean_object* x_109; -x_107 = lean_ctor_get(x_13, 0); -x_108 = lean_ctor_get(x_13, 1); -lean_inc(x_108); +lean_object* x_106; lean_object* x_107; lean_object* x_108; +x_106 = lean_ctor_get(x_13, 0); +x_107 = lean_ctor_get(x_13, 1); lean_inc(x_107); +lean_inc(x_106); lean_dec(x_13); -x_109 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_109, 0, x_107); -lean_ctor_set(x_109, 1, x_108); -return x_109; +x_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_110; +uint8_t x_109; lean_dec(x_8); lean_dec(x_6); lean_dec(x_5); @@ -3607,23 +3609,23 @@ lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_110 = !lean_is_exclusive(x_10); -if (x_110 == 0) +x_109 = !lean_is_exclusive(x_10); +if (x_109 == 0) { return x_10; } else { -lean_object* x_111; lean_object* x_112; lean_object* x_113; -x_111 = lean_ctor_get(x_10, 0); -x_112 = lean_ctor_get(x_10, 1); -lean_inc(x_112); +lean_object* x_110; lean_object* x_111; lean_object* x_112; +x_110 = lean_ctor_get(x_10, 0); +x_111 = lean_ctor_get(x_10, 1); lean_inc(x_111); +lean_inc(x_110); lean_dec(x_10); -x_113 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_113, 0, x_111); -lean_ctor_set(x_113, 1, x_112); -return x_113; +x_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; } } } @@ -4656,103 +4658,100 @@ return x_2; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseCmd___lambda__8(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; 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_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; x_8 = lean_ctor_get(x_1, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); -x_10 = lean_ctor_get(x_9, 1); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_9); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_string_utf8_byte_size(x_11); -lean_dec(x_11); +x_11 = lean_string_utf8_byte_size(x_10); +lean_dec(x_10); lean_inc(x_8); -x_13 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_13, 0, x_8); -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_ctor_get(x_3, 0); +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_8); +lean_ctor_set(x_12, 1, x_11); +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = lean_ctor_get(x_3, 0); +lean_inc(x_14); +x_15 = lean_ctor_get(x_3, 2); lean_inc(x_15); -x_16 = lean_ctor_get(x_3, 2); -lean_inc(x_16); -x_17 = l_Lean_Elab_Command_instInhabitedScope; -x_18 = l_List_head_x21___rarg(x_17, x_16); -lean_dec(x_16); -x_19 = lean_ctor_get(x_18, 1); +x_16 = l_Lean_Elab_Command_instInhabitedScope; +x_17 = l_List_head_x21___rarg(x_16, x_15); +lean_dec(x_15); +x_18 = lean_ctor_get(x_17, 1); +lean_inc(x_18); +x_19 = lean_ctor_get(x_17, 2); lean_inc(x_19); -x_20 = lean_ctor_get(x_18, 2); +x_20 = lean_ctor_get(x_17, 3); lean_inc(x_20); -x_21 = lean_ctor_get(x_18, 3); -lean_inc(x_21); -lean_dec(x_18); -x_22 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_22, 0, x_15); -lean_ctor_set(x_22, 1, x_19); -lean_ctor_set(x_22, 2, x_20); -lean_ctor_set(x_22, 3, x_21); -x_23 = l_Lean_MessageLog_empty; -x_24 = l_Lean_Parser_parseCommand(x_10, x_22, x_1, x_23); -x_25 = lean_ctor_get(x_24, 1); +lean_dec(x_17); +x_21 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_21, 0, x_14); +lean_ctor_set(x_21, 1, x_18); +lean_ctor_set(x_21, 2, x_19); +lean_ctor_set(x_21, 3, x_20); +x_22 = l_Lean_MessageLog_empty; +x_23 = l_Lean_Parser_parseCommand(x_9, x_21, x_1, x_22); +x_24 = lean_ctor_get(x_23, 1); +lean_inc(x_24); +x_25 = lean_ctor_get(x_23, 0); lean_inc(x_25); +lean_dec(x_23); x_26 = lean_ctor_get(x_24, 0); lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_ctor_get(x_25, 0); +x_27 = lean_ctor_get(x_24, 1); lean_inc(x_27); -x_28 = lean_ctor_get(x_25, 1); -lean_inc(x_28); -lean_dec(x_25); +lean_dec(x_24); lean_inc(x_4); -lean_inc(x_27); -lean_inc(x_2); lean_inc(x_26); -x_29 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__3), 9, 7); -lean_closure_set(x_29, 0, x_28); -lean_closure_set(x_29, 1, x_26); -lean_closure_set(x_29, 2, x_3); -lean_closure_set(x_29, 3, x_8); -lean_closure_set(x_29, 4, x_2); -lean_closure_set(x_29, 5, x_27); -lean_closure_set(x_29, 6, x_4); +lean_inc(x_2); +lean_inc(x_25); +x_28 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__3), 9, 7); +lean_closure_set(x_28, 0, x_27); +lean_closure_set(x_28, 1, x_25); +lean_closure_set(x_28, 2, x_3); +lean_closure_set(x_28, 3, x_8); +lean_closure_set(x_28, 4, x_2); +lean_closure_set(x_28, 5, x_26); +lean_closure_set(x_28, 6, x_4); if (lean_obj_tag(x_4) == 0) { -lean_object* x_30; lean_object* x_31; lean_object* x_32; -lean_dec(x_27); +lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_dec(x_26); +lean_dec(x_25); lean_dec(x_2); -x_30 = l_Lean_Language_Lean_process_parseCmd___lambda__8___closed__1; -x_31 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_31, 0, x_30); -lean_closure_set(x_31, 1, x_29); -x_32 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_14, x_31, x_7); -return x_32; +x_29 = l_Lean_Language_Lean_process_parseCmd___lambda__8___closed__1; +x_30 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_30, 0, x_29); +lean_closure_set(x_30, 1, x_28); +x_31 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_30, x_7); +return x_31; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -x_33 = lean_ctor_get(x_4, 0); -lean_inc(x_33); +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_32 = lean_ctor_get(x_4, 0); +lean_inc(x_32); lean_dec(x_4); -x_34 = lean_ctor_get(x_27, 0); -lean_inc(x_34); -lean_dec(x_27); +x_33 = lean_ctor_get(x_26, 0); +lean_inc(x_33); +lean_dec(x_26); lean_inc(x_2); -x_35 = lean_alloc_closure((void*)(l_Lean_Language_Lean_isBeforeEditPos___boxed), 3, 2); -lean_closure_set(x_35, 0, x_34); -lean_closure_set(x_35, 1, x_2); -x_36 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7___boxed), 6, 4); -lean_closure_set(x_36, 0, x_33); -lean_closure_set(x_36, 1, x_29); -lean_closure_set(x_36, 2, x_26); -lean_closure_set(x_36, 3, x_2); -x_37 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); -lean_closure_set(x_37, 0, x_35); -lean_closure_set(x_37, 1, x_36); -x_38 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_14, x_37, x_7); -return x_38; +x_34 = lean_alloc_closure((void*)(l_Lean_Language_Lean_isBeforeEditPos___boxed), 3, 2); +lean_closure_set(x_34, 0, x_33); +lean_closure_set(x_34, 1, x_2); +x_35 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseCmd___lambda__7___boxed), 6, 4); +lean_closure_set(x_35, 0, x_32); +lean_closure_set(x_35, 1, x_28); +lean_closure_set(x_35, 2, x_25); +lean_closure_set(x_35, 3, x_2); +x_36 = lean_alloc_closure((void*)(l_EStateM_bind___rarg), 3, 2); +lean_closure_set(x_36, 0, x_34); +lean_closure_set(x_36, 1, x_35); +x_37 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_13, x_36, x_7); +return x_37; } } } @@ -5577,14 +5576,7 @@ x_6 = lean_apply_3(x_1, x_2, x_5, x_4); return x_6; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_inc(x_3); -return x_3; -} -} -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5594,7 +5586,7 @@ x_3 = lean_nat_add(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2() { _start: { lean_object* x_1; @@ -5602,7 +5594,7 @@ x_1 = l_Lean_maxRecDepth; return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5614,7 +5606,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5626,7 +5618,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -5638,14 +5630,14 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___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 = lean_unsigned_to_nat(0u); -x_2 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4; -x_3 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; -x_4 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5; +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4; +x_3 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; +x_4 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5; x_5 = lean_alloc_ctor(0, 9, 0); lean_ctor_set(x_5, 0, x_1); lean_ctor_set(x_5, 1, x_1); @@ -5659,7 +5651,7 @@ lean_ctor_set(x_5, 8, x_3); return x_5; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7() { _start: { lean_object* x_1; @@ -5667,21 +5659,21 @@ x_1 = lean_mk_string_from_bytes("_import", 7); return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7; +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8; +x_1 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8; x_2 = lean_unsigned_to_nat(1u); x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -5689,7 +5681,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10() { _start: { lean_object* x_1; @@ -5697,17 +5689,17 @@ x_1 = lean_mk_string_from_bytes("header", 6); return x_1; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10; +x_2 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12() { +static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12() { _start: { lean_object* x_1; lean_object* x_2; @@ -5716,7 +5708,7 @@ x_2 = lean_mk_empty_array_with_capacity(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_Language_Lean_process_processHeader___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; lean_object* x_13; lean_object* x_14; uint8_t x_15; @@ -5742,14 +5734,14 @@ x_19 = lean_ctor_get(x_14, 1); lean_dec(x_19); x_20 = lean_ctor_get(x_14, 0); lean_dec(x_20); -x_21 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; +x_21 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; x_22 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_21); lean_dec(x_4); x_23 = lean_ctor_get(x_5, 2); x_24 = lean_box(0); -x_25 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6; +x_25 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; x_26 = lean_box(0); -x_27 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9; +x_27 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; lean_inc(x_23); lean_inc(x_13); x_28 = lean_alloc_ctor(0, 7, 0); @@ -5762,7 +5754,7 @@ lean_ctor_set(x_28, 5, x_24); lean_ctor_set(x_28, 6, x_27); x_29 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_29, 0, x_28); -x_30 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_30 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11; lean_inc(x_6); x_31 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_31, 0, x_30); @@ -5783,18 +5775,18 @@ lean_ctor_set(x_39, 1, x_38); x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_29); lean_ctor_set(x_40, 1, x_39); -x_41 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12; +x_41 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; x_42 = lean_array_push(x_41, x_40); x_43 = l_Array_toPArray_x27___rarg(x_42); lean_dec(x_42); x_44 = 1; -x_45 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; +x_45 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; lean_inc(x_43); x_46 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_46, 0, x_45); lean_ctor_set(x_46, 1, x_43); lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_44); -x_47 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; +x_47 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; lean_ctor_set(x_14, 6, x_46); lean_ctor_set(x_14, 4, x_22); lean_ctor_set(x_14, 3, x_47); @@ -5927,14 +5919,14 @@ lean_inc(x_92); lean_inc(x_91); lean_inc(x_90); lean_dec(x_14); -x_93 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2; +x_93 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2; x_94 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_4, x_93); lean_dec(x_4); x_95 = lean_ctor_get(x_5, 2); x_96 = lean_box(0); -x_97 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6; +x_97 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6; x_98 = lean_box(0); -x_99 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9; +x_99 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9; lean_inc(x_95); lean_inc(x_13); x_100 = lean_alloc_ctor(0, 7, 0); @@ -5947,7 +5939,7 @@ lean_ctor_set(x_100, 5, x_96); lean_ctor_set(x_100, 6, x_99); x_101 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_101, 0, x_100); -x_102 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11; +x_102 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11; lean_inc(x_6); x_103 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_103, 0, x_102); @@ -5968,18 +5960,18 @@ lean_ctor_set(x_111, 1, x_110); x_112 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_112, 0, x_101); lean_ctor_set(x_112, 1, x_111); -x_113 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12; +x_113 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12; x_114 = lean_array_push(x_113, x_112); x_115 = l_Array_toPArray_x27___rarg(x_114); lean_dec(x_114); x_116 = 1; -x_117 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3; +x_117 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3; lean_inc(x_115); x_118 = lean_alloc_ctor(0, 2, 1); lean_ctor_set(x_118, 0, x_117); lean_ctor_set(x_118, 1, x_115); lean_ctor_set_uint8(x_118, sizeof(void*)*2, x_116); -x_119 = l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1; +x_119 = l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1; x_120 = lean_alloc_ctor(0, 8, 0); lean_ctor_set(x_120, 0, x_13); lean_ctor_set(x_120, 1, x_3); @@ -6066,186 +6058,172 @@ return x_144; } } } -static lean_object* _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1() { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__3___boxed), 3, 0); -return x_1; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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) { -_start: +lean_object* x_7; uint32_t x_8; lean_object* x_9; uint8_t x_10; lean_object* x_11; +x_7 = lean_ctor_get(x_4, 1); +lean_inc(x_7); +x_8 = lean_ctor_get_uint32(x_4, sizeof(void*)*2); +x_9 = l_Lean_MessageLog_empty; +x_10 = 1; +lean_inc(x_7); +x_11 = l_Lean_Elab_processHeader(x_1, x_7, x_9, x_2, x_8, x_10, x_6); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; uint32_t x_12; lean_object* x_13; uint8_t x_14; lean_object* x_15; -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); -x_10 = l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1; -x_11 = l_List_forIn_loop___at_Lean_KVMap_mergeBy___spec__1(x_10, x_5, x_9); -x_12 = lean_ctor_get_uint32(x_8, sizeof(void*)*2); -x_13 = l_Lean_MessageLog_empty; -x_14 = 1; -lean_inc(x_11); -x_15 = l_Lean_Elab_processHeader(x_2, x_11, x_13, x_3, x_12, x_14, x_7); -if (lean_obj_tag(x_15) == 0) +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_12 = lean_ctor_get(x_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); +x_15 = lean_ctor_get(x_12, 1); +lean_inc(x_15); +lean_dec(x_12); +lean_inc(x_15); +x_16 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_15, x_13); +x_17 = !lean_is_exclusive(x_16); +if (x_17 == 0) { -lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; uint8_t x_21; -x_16 = 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_object* x_18; lean_object* x_19; uint8_t x_20; x_18 = lean_ctor_get(x_16, 0); -lean_inc(x_18); x_19 = lean_ctor_get(x_16, 1); -lean_inc(x_19); -lean_dec(x_16); -lean_inc(x_19); -x_20 = l_Lean_Language_Snapshot_Diagnostics_ofMessageLog(x_19, x_17); -x_21 = !lean_is_exclusive(x_20); -if (x_21 == 0) -{ -lean_object* x_22; lean_object* x_23; uint8_t x_24; -x_22 = lean_ctor_get(x_20, 0); -x_23 = lean_ctor_get(x_20, 1); -lean_inc(x_19); -x_24 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_19); -if (x_24 == 0) +lean_inc(x_15); +x_20 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_15); +if (x_20 == 0) { -lean_object* x_25; lean_object* x_26; -lean_free_object(x_20); -x_25 = lean_box(0); -x_26 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_8, x_18, x_19, x_11, x_3, x_2, x_4, x_22, x_25, x_6, x_23); -lean_dec(x_3); -return x_26; +lean_object* x_21; lean_object* x_22; +lean_free_object(x_16); +x_21 = lean_box(0); +x_22 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_4, x_14, x_15, x_7, x_2, x_1, x_3, x_18, x_21, x_5, x_19); +lean_dec(x_2); +return x_22; } else { -lean_object* x_27; lean_object* x_28; lean_object* x_29; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_23; lean_object* x_24; lean_object* x_25; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_27 = lean_box(0); -x_28 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_28, 0, x_22); -lean_ctor_set(x_28, 1, x_27); -lean_ctor_set_uint8(x_28, sizeof(void*)*2, x_14); -x_29 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_29, 0, x_28); -lean_ctor_set(x_29, 1, x_27); -lean_ctor_set(x_20, 0, x_29); -return x_20; +lean_dec(x_1); +x_23 = lean_box(0); +x_24 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_24, 0, x_18); +lean_ctor_set(x_24, 1, x_23); +lean_ctor_set_uint8(x_24, sizeof(void*)*2, x_10); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +lean_ctor_set(x_16, 0, x_25); +return x_16; } } else { -lean_object* x_30; lean_object* x_31; uint8_t 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); -lean_inc(x_19); -x_32 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_19); -if (x_32 == 0) +lean_object* x_26; lean_object* x_27; uint8_t x_28; +x_26 = lean_ctor_get(x_16, 0); +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_inc(x_26); +lean_dec(x_16); +lean_inc(x_15); +x_28 = l_Lean_PersistentArray_anyM___at_Lean_MessageLog_hasErrors___spec__1(x_15); +if (x_28 == 0) { -lean_object* x_33; lean_object* x_34; -x_33 = lean_box(0); -x_34 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_8, x_18, x_19, x_11, x_3, x_2, x_4, x_30, x_33, x_6, x_31); -lean_dec(x_3); -return x_34; +lean_object* x_29; lean_object* x_30; +x_29 = lean_box(0); +x_30 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_4, x_14, x_15, x_7, x_2, x_1, x_3, x_26, x_29, x_5, x_27); +lean_dec(x_2); +return x_30; } else { -lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; -lean_dec(x_19); -lean_dec(x_18); -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_6); +lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; +lean_dec(x_15); +lean_dec(x_14); +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_35 = lean_box(0); -x_36 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_36, 0, x_30); -lean_ctor_set(x_36, 1, x_35); -lean_ctor_set_uint8(x_36, sizeof(void*)*2, x_14); -x_37 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_37, 0, x_36); -lean_ctor_set(x_37, 1, 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_31); -return x_38; +lean_dec(x_1); +x_31 = lean_box(0); +x_32 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_32, 0, x_26); +lean_ctor_set(x_32, 1, x_31); +lean_ctor_set_uint8(x_32, sizeof(void*)*2, x_10); +x_33 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_33, 0, x_32); +lean_ctor_set(x_33, 1, x_31); +x_34 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_34, 0, x_33); +lean_ctor_set(x_34, 1, x_27); +return x_34; } } } else { -uint8_t x_39; -lean_dec(x_11); -lean_dec(x_8); -lean_dec(x_6); +uint8_t x_35; +lean_dec(x_7); +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_39 = !lean_is_exclusive(x_15); -if (x_39 == 0) +lean_dec(x_1); +x_35 = !lean_is_exclusive(x_11); +if (x_35 == 0) { -return x_15; +return x_11; } else { -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); -lean_dec(x_15); -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_36; lean_object* x_37; lean_object* x_38; +x_36 = lean_ctor_get(x_11, 0); +x_37 = lean_ctor_get(x_11, 1); +lean_inc(x_37); +lean_inc(x_36); +lean_dec(x_11); +x_38 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_38, 0, x_36); +lean_ctor_set(x_38, 1, x_37); +return x_38; } } } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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) { _start: { -if (lean_obj_tag(x_5) == 0) +if (lean_obj_tag(x_4) == 0) { -lean_object* x_8; lean_object* x_9; -lean_dec(x_6); -lean_dec(x_4); +lean_object* x_7; lean_object* x_8; +lean_dec(x_5); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_8 = lean_ctor_get(x_5, 0); -lean_inc(x_8); -lean_dec(x_5); -x_9 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_9, 0, x_8); -lean_ctor_set(x_9, 1, x_7); -return x_9; +x_7 = lean_ctor_get(x_4, 0); +lean_inc(x_7); +lean_dec(x_4); +x_8 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_8, 0, x_7); +lean_ctor_set(x_8, 1, x_6); +return x_8; } else { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_5, 0); -lean_inc(x_10); -lean_dec(x_5); -x_11 = l_Lean_Language_Lean_process_processHeader___lambda__5(x_1, x_2, x_3, x_4, x_10, x_6, x_7); -return x_11; +lean_object* x_9; lean_object* x_10; +x_9 = lean_ctor_get(x_4, 0); +lean_inc(x_9); +lean_dec(x_4); +x_10 = l_Lean_Language_Lean_process_processHeader___lambda__4(x_1, x_2, x_3, x_9, x_5, x_6); +return x_10; } } } @@ -6260,58 +6238,44 @@ return x_1; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { -lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; +lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_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_6 = lean_ctor_get(x_4, 0); lean_inc(x_6); -x_7 = lean_ctor_get(x_6, 1); +x_7 = lean_ctor_get(x_6, 0); lean_inc(x_7); -x_8 = lean_ctor_get(x_7, 0); -lean_inc(x_8); -x_9 = lean_string_utf8_byte_size(x_8); -lean_dec(x_8); -x_10 = lean_unsigned_to_nat(0u); -x_11 = lean_alloc_ctor(0, 2, 0); +x_8 = lean_string_utf8_byte_size(x_7); +lean_dec(x_7); +x_9 = lean_unsigned_to_nat(0u); +x_10 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_10, 0, x_9); +lean_ctor_set(x_10, 1, x_8); +x_11 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_11, 1, x_9); -x_12 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_12, 0, x_11); lean_inc(x_2); -x_13 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__2), 4, 2); -lean_closure_set(x_13, 0, x_1); -lean_closure_set(x_13, 1, x_2); -x_14 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__6), 7, 4); -lean_closure_set(x_14, 0, x_6); -lean_closure_set(x_14, 1, x_2); -lean_closure_set(x_14, 2, x_7); -lean_closure_set(x_14, 3, x_3); -x_15 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg), 4, 2); -lean_closure_set(x_15, 0, x_13); -lean_closure_set(x_15, 1, x_14); -x_16 = l_Lean_Language_Lean_process_processHeader___closed__1; -x_17 = lean_alloc_closure((void*)(l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg), 4, 3); -lean_closure_set(x_17, 0, x_16); -lean_closure_set(x_17, 1, x_15); -lean_closure_set(x_17, 2, x_4); -x_18 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_12, x_17, x_5); -return x_18; -} -} -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___lambda__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { -_start: -{ -lean_object* x_4; -x_4 = l_Lean_Language_Lean_process_processHeader___lambda__3(x_1, x_2, x_3); -lean_dec(x_3); -lean_dec(x_2); -lean_dec(x_1); -return x_4; +x_12 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__2), 4, 2); +lean_closure_set(x_12, 0, x_1); +lean_closure_set(x_12, 1, x_2); +x_13 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_processHeader___lambda__5), 6, 3); +lean_closure_set(x_13, 0, x_2); +lean_closure_set(x_13, 1, x_6); +lean_closure_set(x_13, 2, x_3); +x_14 = lean_alloc_closure((void*)(l_ReaderT_bind___at_Lean_Language_Lean_process_processHeader___spec__2___rarg), 4, 2); +lean_closure_set(x_14, 0, x_12); +lean_closure_set(x_14, 1, x_13); +x_15 = l_Lean_Language_Lean_process_processHeader___closed__1; +x_16 = lean_alloc_closure((void*)(l___private_Lean_Language_Lean_0__Lean_Language_Lean_withHeaderExceptions___rarg), 4, 3); +lean_closure_set(x_16, 0, x_15); +lean_closure_set(x_16, 1, x_14); +lean_closure_set(x_16, 2, x_4); +x_17 = l_Lean_Language_SnapshotTask_ofIO___rarg(x_11, x_16, x_5); +return x_17; } } -LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_processHeader___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_Language_Lean_process_processHeader___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_Language_Lean_process_processHeader___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_Language_Lean_process_processHeader___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_9); lean_dec(x_5); return x_12; @@ -7238,545 +7202,540 @@ return x_12; LEAN_EXPORT lean_object* l_Lean_Language_Lean_process_parseHeader(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; -x_5 = lean_ctor_get(x_3, 0); -lean_inc(x_5); if (lean_obj_tag(x_2) == 0) { -lean_object* x_6; lean_object* x_7; lean_object* x_8; -x_6 = lean_ctor_get(x_5, 1); -lean_inc(x_6); -lean_dec(x_5); -x_7 = lean_box(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); +x_6 = lean_box(0); lean_inc(x_3); -x_8 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_6, x_1, x_2, x_3, x_7, x_3, x_4); -return x_8; +x_7 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_5, x_1, x_2, x_3, x_6, x_3, x_4); +return x_7; } else { -lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; -x_9 = lean_ctor_get(x_5, 1); +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; +x_8 = lean_ctor_get(x_3, 0); +lean_inc(x_8); +x_9 = lean_ctor_get(x_2, 0); lean_inc(x_9); -lean_dec(x_5); -x_10 = lean_ctor_get(x_2, 0); -lean_inc(x_10); -lean_inc(x_10); -x_11 = l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult(x_10); -x_12 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_11, x_4); +lean_inc(x_9); +x_10 = l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult(x_9); +x_11 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_10, x_4); +if (lean_obj_tag(x_11) == 0) +{ +lean_object* x_12; +x_12 = lean_ctor_get(x_11, 0); +lean_inc(x_12); if (lean_obj_tag(x_12) == 0) { -lean_object* x_13; -x_13 = lean_ctor_get(x_12, 0); +lean_object* x_13; lean_object* x_14; lean_object* x_15; +lean_dec(x_9); +x_13 = lean_ctor_get(x_11, 1); lean_inc(x_13); -if (lean_obj_tag(x_13) == 0) -{ -lean_object* x_14; lean_object* x_15; lean_object* x_16; -lean_dec(x_10); -x_14 = lean_ctor_get(x_12, 1); -lean_inc(x_14); -lean_dec(x_12); -x_15 = lean_box(0); +lean_dec(x_11); +x_14 = lean_box(0); lean_inc(x_3); -x_16 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_9, x_1, x_2, x_3, x_15, x_3, x_14); -return x_16; +x_15 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_8, x_1, x_2, x_3, x_14, x_3, x_13); +return x_15; } else { -lean_object* x_17; -x_17 = lean_ctor_get(x_13, 0); -lean_inc(x_17); -lean_dec(x_13); -if (lean_obj_tag(x_17) == 0) -{ -lean_object* x_18; lean_object* x_19; lean_object* x_20; -lean_dec(x_10); -x_18 = lean_ctor_get(x_12, 1); -lean_inc(x_18); +lean_object* x_16; +x_16 = lean_ctor_get(x_12, 0); +lean_inc(x_16); lean_dec(x_12); -x_19 = lean_box(0); +if (lean_obj_tag(x_16) == 0) +{ +lean_object* x_17; lean_object* x_18; lean_object* x_19; +lean_dec(x_9); +x_17 = lean_ctor_get(x_11, 1); +lean_inc(x_17); +lean_dec(x_11); +x_18 = lean_box(0); lean_inc(x_3); -x_20 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_9, x_1, x_2, x_3, x_19, x_3, x_18); -return x_20; +x_19 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_8, x_1, x_2, x_3, x_18, x_3, x_17); +return x_19; } else { -lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_21 = lean_ctor_get(x_12, 1); +lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_20 = lean_ctor_get(x_11, 1); +lean_inc(x_20); +lean_dec(x_11); +x_21 = lean_ctor_get(x_16, 0); lean_inc(x_21); -lean_dec(x_12); -x_22 = lean_ctor_get(x_17, 0); +lean_dec(x_16); +x_22 = lean_ctor_get(x_21, 1); lean_inc(x_22); -lean_dec(x_17); -x_23 = lean_ctor_get(x_22, 1); -lean_inc(x_23); -lean_dec(x_22); -x_24 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_23, x_21); +lean_dec(x_21); +x_23 = l_Lean_Language_SnapshotTask_get_x3f___rarg(x_22, x_20); +if (lean_obj_tag(x_23) == 0) +{ +lean_object* x_24; +x_24 = lean_ctor_get(x_23, 0); +lean_inc(x_24); if (lean_obj_tag(x_24) == 0) { -lean_object* x_25; -x_25 = lean_ctor_get(x_24, 0); +lean_object* x_25; lean_object* x_26; lean_object* x_27; +lean_dec(x_9); +x_25 = lean_ctor_get(x_23, 1); lean_inc(x_25); -if (lean_obj_tag(x_25) == 0) -{ -lean_object* x_26; lean_object* x_27; lean_object* x_28; -lean_dec(x_10); -x_26 = lean_ctor_get(x_24, 1); -lean_inc(x_26); -lean_dec(x_24); -x_27 = lean_box(0); +lean_dec(x_23); +x_26 = lean_box(0); lean_inc(x_3); -x_28 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_9, x_1, x_2, x_3, x_27, x_3, x_26); -return x_28; +x_27 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_8, x_1, x_2, x_3, x_26, x_3, x_25); +return x_27; } else { -lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint8_t x_36; -x_29 = lean_ctor_get(x_25, 0); +lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; uint8_t x_35; +x_28 = lean_ctor_get(x_24, 0); +lean_inc(x_28); +lean_dec(x_24); +x_29 = lean_ctor_get(x_23, 1); lean_inc(x_29); -lean_dec(x_25); -x_30 = lean_ctor_get(x_24, 1); +lean_dec(x_23); +x_30 = lean_ctor_get(x_28, 0); lean_inc(x_30); -lean_dec(x_24); -x_31 = lean_ctor_get(x_29, 0); +lean_dec(x_28); +x_31 = lean_ctor_get(x_30, 2); lean_inc(x_31); -lean_dec(x_29); -x_32 = lean_ctor_get(x_31, 2); +lean_dec(x_30); +x_32 = lean_ctor_get(x_31, 0); lean_inc(x_32); lean_dec(x_31); -x_33 = lean_ctor_get(x_32, 0); -lean_inc(x_33); +x_33 = l_Lean_Language_Lean_isBeforeEditPos(x_32, x_3, x_29); lean_dec(x_32); -x_34 = l_Lean_Language_Lean_isBeforeEditPos(x_33, x_3, x_30); -lean_dec(x_33); -x_35 = lean_ctor_get(x_34, 0); -lean_inc(x_35); -x_36 = lean_unbox(x_35); -lean_dec(x_35); -if (x_36 == 0) -{ -lean_object* x_37; lean_object* x_38; lean_object* x_39; -lean_dec(x_10); -x_37 = lean_ctor_get(x_34, 1); -lean_inc(x_37); +x_34 = lean_ctor_get(x_33, 0); +lean_inc(x_34); +x_35 = lean_unbox(x_34); lean_dec(x_34); -x_38 = lean_box(0); +if (x_35 == 0) +{ +lean_object* x_36; lean_object* x_37; lean_object* x_38; +lean_dec(x_9); +x_36 = lean_ctor_get(x_33, 1); +lean_inc(x_36); +lean_dec(x_33); +x_37 = lean_box(0); lean_inc(x_3); -x_39 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_9, x_1, x_2, x_3, x_38, x_3, x_37); -return x_39; +x_38 = l_Lean_Language_Lean_process_parseHeader___lambda__10(x_8, x_1, x_2, x_3, x_37, x_3, x_36); +return x_38; } else { -lean_object* x_40; +lean_object* x_39; lean_dec(x_2); lean_dec(x_1); -x_40 = lean_ctor_get(x_10, 3); -lean_inc(x_40); -if (lean_obj_tag(x_40) == 0) +x_39 = lean_ctor_get(x_9, 3); +lean_inc(x_39); +if (lean_obj_tag(x_39) == 0) { -uint8_t x_41; -lean_dec(x_9); +uint8_t x_40; +lean_dec(x_8); lean_dec(x_3); -x_41 = !lean_is_exclusive(x_34); -if (x_41 == 0) +x_40 = !lean_is_exclusive(x_33); +if (x_40 == 0) { -lean_object* x_42; -x_42 = lean_ctor_get(x_34, 0); -lean_dec(x_42); -lean_ctor_set(x_34, 0, x_10); -return x_34; +lean_object* x_41; +x_41 = lean_ctor_get(x_33, 0); +lean_dec(x_41); +lean_ctor_set(x_33, 0, x_9); +return x_33; } else { -lean_object* x_43; lean_object* x_44; -x_43 = lean_ctor_get(x_34, 1); -lean_inc(x_43); -lean_dec(x_34); -x_44 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_44, 0, x_10); -lean_ctor_set(x_44, 1, x_43); -return x_44; +lean_object* x_42; lean_object* x_43; +x_42 = lean_ctor_get(x_33, 1); +lean_inc(x_42); +lean_dec(x_33); +x_43 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_43, 0, x_9); +lean_ctor_set(x_43, 1, x_42); +return x_43; } } else { -uint8_t x_45; -x_45 = !lean_is_exclusive(x_40); -if (x_45 == 0) +uint8_t x_44; +x_44 = !lean_is_exclusive(x_39); +if (x_44 == 0) { -lean_object* x_46; lean_object* x_47; uint8_t x_48; -x_46 = lean_ctor_get(x_40, 0); -x_47 = lean_ctor_get(x_34, 1); -lean_inc(x_47); -lean_dec(x_34); -x_48 = !lean_is_exclusive(x_10); -if (x_48 == 0) +lean_object* x_45; lean_object* x_46; uint8_t x_47; +x_45 = lean_ctor_get(x_39, 0); +x_46 = lean_ctor_get(x_33, 1); +lean_inc(x_46); +lean_dec(x_33); +x_47 = !lean_is_exclusive(x_9); +if (x_47 == 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; uint8_t x_56; -x_49 = lean_ctor_get(x_10, 0); -x_50 = lean_ctor_get(x_10, 2); -x_51 = lean_ctor_get(x_10, 3); +lean_object* x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_48 = lean_ctor_get(x_9, 0); +x_49 = lean_ctor_get(x_9, 2); +x_50 = lean_ctor_get(x_9, 3); +lean_dec(x_50); +x_51 = lean_ctor_get(x_9, 1); lean_dec(x_51); -x_52 = lean_ctor_get(x_10, 1); -lean_dec(x_52); -x_53 = lean_ctor_get(x_46, 0); +x_52 = lean_ctor_get(x_45, 0); +lean_inc(x_52); +x_53 = lean_ctor_get(x_45, 1); lean_inc(x_53); -x_54 = lean_ctor_get(x_46, 1); -lean_inc(x_54); -lean_inc(x_46); -x_55 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); -lean_closure_set(x_55, 0, x_46); -lean_closure_set(x_55, 1, x_3); -x_56 = !lean_is_exclusive(x_46); -if (x_56 == 0) +lean_inc(x_45); +x_54 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); +lean_closure_set(x_54, 0, x_45); +lean_closure_set(x_54, 1, x_3); +x_55 = !lean_is_exclusive(x_45); +if (x_55 == 0) { -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_46, 1); +lean_object* x_56; lean_object* x_57; lean_object* x_58; uint8_t x_59; lean_object* x_60; +x_56 = lean_ctor_get(x_45, 1); +lean_dec(x_56); +x_57 = lean_ctor_get(x_45, 0); lean_dec(x_57); -x_58 = lean_ctor_get(x_46, 0); -lean_dec(x_58); -x_59 = lean_ctor_get(x_54, 0); -lean_inc(x_59); -x_60 = 1; -x_61 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_54, x_55, x_59, x_60, x_47); -if (lean_obj_tag(x_61) == 0) -{ -uint8_t x_62; -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) -{ -lean_object* x_63; -x_63 = lean_ctor_get(x_61, 0); -lean_ctor_set(x_46, 1, x_63); -lean_ctor_set(x_10, 1, x_9); -lean_ctor_set(x_61, 0, x_10); -return x_61; +x_58 = lean_ctor_get(x_53, 0); +lean_inc(x_58); +x_59 = 1; +x_60 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_53, x_54, x_58, x_59, x_46); +if (lean_obj_tag(x_60) == 0) +{ +uint8_t x_61; +x_61 = !lean_is_exclusive(x_60); +if (x_61 == 0) +{ +lean_object* x_62; +x_62 = lean_ctor_get(x_60, 0); +lean_ctor_set(x_45, 1, x_62); +lean_ctor_set(x_9, 1, x_8); +lean_ctor_set(x_60, 0, x_9); +return x_60; } else { -lean_object* x_64; lean_object* x_65; lean_object* x_66; -x_64 = lean_ctor_get(x_61, 0); -x_65 = lean_ctor_get(x_61, 1); -lean_inc(x_65); +lean_object* x_63; lean_object* x_64; lean_object* x_65; +x_63 = lean_ctor_get(x_60, 0); +x_64 = lean_ctor_get(x_60, 1); lean_inc(x_64); -lean_dec(x_61); -lean_ctor_set(x_46, 1, x_64); -lean_ctor_set(x_10, 1, x_9); -x_66 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_66, 0, x_10); -lean_ctor_set(x_66, 1, x_65); -return x_66; +lean_inc(x_63); +lean_dec(x_60); +lean_ctor_set(x_45, 1, x_63); +lean_ctor_set(x_9, 1, x_8); +x_65 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_65, 0, x_9); +lean_ctor_set(x_65, 1, x_64); +return x_65; } } else { -uint8_t x_67; -lean_free_object(x_46); -lean_dec(x_53); -lean_free_object(x_10); -lean_dec(x_50); +uint8_t x_66; +lean_free_object(x_45); +lean_dec(x_52); +lean_free_object(x_9); lean_dec(x_49); -lean_free_object(x_40); -lean_dec(x_9); -x_67 = !lean_is_exclusive(x_61); -if (x_67 == 0) +lean_dec(x_48); +lean_free_object(x_39); +lean_dec(x_8); +x_66 = !lean_is_exclusive(x_60); +if (x_66 == 0) { -return x_61; +return x_60; } else { -lean_object* x_68; lean_object* x_69; lean_object* x_70; -x_68 = lean_ctor_get(x_61, 0); -x_69 = lean_ctor_get(x_61, 1); -lean_inc(x_69); +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_dec(x_61); -x_70 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_70, 0, x_68); -lean_ctor_set(x_70, 1, x_69); -return x_70; +lean_inc(x_67); +lean_dec(x_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; } } } else { -lean_object* x_71; uint8_t x_72; lean_object* x_73; -lean_dec(x_46); -x_71 = lean_ctor_get(x_54, 0); -lean_inc(x_71); -x_72 = 1; -x_73 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_54, x_55, x_71, x_72, x_47); -if (lean_obj_tag(x_73) == 0) +lean_object* x_70; uint8_t x_71; lean_object* x_72; +lean_dec(x_45); +x_70 = lean_ctor_get(x_53, 0); +lean_inc(x_70); +x_71 = 1; +x_72 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_53, x_54, x_70, x_71, x_46); +if (lean_obj_tag(x_72) == 0) { -lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; lean_object* x_78; -x_74 = lean_ctor_get(x_73, 0); +lean_object* x_73; lean_object* x_74; lean_object* x_75; lean_object* x_76; lean_object* x_77; +x_73 = lean_ctor_get(x_72, 0); +lean_inc(x_73); +x_74 = lean_ctor_get(x_72, 1); lean_inc(x_74); -x_75 = lean_ctor_get(x_73, 1); -lean_inc(x_75); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_76 = x_73; +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_75 = x_72; } else { - lean_dec_ref(x_73); - x_76 = lean_box(0); + lean_dec_ref(x_72); + x_75 = lean_box(0); } -x_77 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_77, 0, x_53); -lean_ctor_set(x_77, 1, x_74); -lean_ctor_set(x_40, 0, x_77); -lean_ctor_set(x_10, 1, x_9); -if (lean_is_scalar(x_76)) { - x_78 = lean_alloc_ctor(0, 2, 0); +x_76 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_76, 0, x_52); +lean_ctor_set(x_76, 1, x_73); +lean_ctor_set(x_39, 0, x_76); +lean_ctor_set(x_9, 1, x_8); +if (lean_is_scalar(x_75)) { + x_77 = lean_alloc_ctor(0, 2, 0); } else { - x_78 = x_76; + x_77 = x_75; } -lean_ctor_set(x_78, 0, x_10); -lean_ctor_set(x_78, 1, x_75); -return x_78; +lean_ctor_set(x_77, 0, x_9); +lean_ctor_set(x_77, 1, x_74); +return x_77; } else { -lean_object* x_79; lean_object* x_80; lean_object* x_81; lean_object* x_82; -lean_dec(x_53); -lean_free_object(x_10); -lean_dec(x_50); +lean_object* x_78; lean_object* x_79; lean_object* x_80; lean_object* x_81; +lean_dec(x_52); +lean_free_object(x_9); lean_dec(x_49); -lean_free_object(x_40); -lean_dec(x_9); -x_79 = lean_ctor_get(x_73, 0); +lean_dec(x_48); +lean_free_object(x_39); +lean_dec(x_8); +x_78 = lean_ctor_get(x_72, 0); +lean_inc(x_78); +x_79 = lean_ctor_get(x_72, 1); lean_inc(x_79); -x_80 = lean_ctor_get(x_73, 1); -lean_inc(x_80); -if (lean_is_exclusive(x_73)) { - lean_ctor_release(x_73, 0); - lean_ctor_release(x_73, 1); - x_81 = x_73; +if (lean_is_exclusive(x_72)) { + lean_ctor_release(x_72, 0); + lean_ctor_release(x_72, 1); + x_80 = x_72; } else { - lean_dec_ref(x_73); - x_81 = lean_box(0); + lean_dec_ref(x_72); + x_80 = lean_box(0); } -if (lean_is_scalar(x_81)) { - x_82 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_80)) { + x_81 = lean_alloc_ctor(1, 2, 0); } else { - x_82 = x_81; + x_81 = x_80; } -lean_ctor_set(x_82, 0, x_79); -lean_ctor_set(x_82, 1, x_80); -return x_82; +lean_ctor_set(x_81, 0, x_78); +lean_ctor_set(x_81, 1, x_79); +return x_81; } } } 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; uint8_t x_90; lean_object* x_91; -x_83 = lean_ctor_get(x_10, 0); -x_84 = lean_ctor_get(x_10, 2); -lean_inc(x_84); +lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; uint8_t x_89; lean_object* x_90; +x_82 = lean_ctor_get(x_9, 0); +x_83 = lean_ctor_get(x_9, 2); lean_inc(x_83); -lean_dec(x_10); -x_85 = lean_ctor_get(x_46, 0); +lean_inc(x_82); +lean_dec(x_9); +x_84 = lean_ctor_get(x_45, 0); +lean_inc(x_84); +x_85 = lean_ctor_get(x_45, 1); lean_inc(x_85); -x_86 = lean_ctor_get(x_46, 1); -lean_inc(x_86); -lean_inc(x_46); -x_87 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); -lean_closure_set(x_87, 0, x_46); -lean_closure_set(x_87, 1, x_3); -if (lean_is_exclusive(x_46)) { - lean_ctor_release(x_46, 0); - lean_ctor_release(x_46, 1); - x_88 = x_46; +lean_inc(x_45); +x_86 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); +lean_closure_set(x_86, 0, x_45); +lean_closure_set(x_86, 1, x_3); +if (lean_is_exclusive(x_45)) { + lean_ctor_release(x_45, 0); + lean_ctor_release(x_45, 1); + x_87 = x_45; } else { - lean_dec_ref(x_46); - x_88 = lean_box(0); + lean_dec_ref(x_45); + x_87 = lean_box(0); } -x_89 = lean_ctor_get(x_86, 0); -lean_inc(x_89); -x_90 = 1; -x_91 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_86, x_87, x_89, x_90, x_47); -if (lean_obj_tag(x_91) == 0) +x_88 = lean_ctor_get(x_85, 0); +lean_inc(x_88); +x_89 = 1; +x_90 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_85, x_86, x_88, x_89, x_46); +if (lean_obj_tag(x_90) == 0) { -lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; -x_92 = lean_ctor_get(x_91, 0); +lean_object* x_91; lean_object* x_92; lean_object* x_93; lean_object* x_94; lean_object* x_95; lean_object* x_96; +x_91 = lean_ctor_get(x_90, 0); +lean_inc(x_91); +x_92 = lean_ctor_get(x_90, 1); lean_inc(x_92); -x_93 = lean_ctor_get(x_91, 1); -lean_inc(x_93); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_94 = x_91; +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_93 = x_90; } else { - lean_dec_ref(x_91); - x_94 = lean_box(0); + lean_dec_ref(x_90); + x_93 = lean_box(0); } -if (lean_is_scalar(x_88)) { - x_95 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_87)) { + x_94 = lean_alloc_ctor(0, 2, 0); } else { - x_95 = x_88; -} -lean_ctor_set(x_95, 0, x_85); -lean_ctor_set(x_95, 1, x_92); -lean_ctor_set(x_40, 0, x_95); -x_96 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_96, 0, x_83); -lean_ctor_set(x_96, 1, x_9); -lean_ctor_set(x_96, 2, x_84); -lean_ctor_set(x_96, 3, x_40); -if (lean_is_scalar(x_94)) { - x_97 = lean_alloc_ctor(0, 2, 0); + x_94 = x_87; +} +lean_ctor_set(x_94, 0, x_84); +lean_ctor_set(x_94, 1, x_91); +lean_ctor_set(x_39, 0, x_94); +x_95 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_95, 0, x_82); +lean_ctor_set(x_95, 1, x_8); +lean_ctor_set(x_95, 2, x_83); +lean_ctor_set(x_95, 3, x_39); +if (lean_is_scalar(x_93)) { + x_96 = lean_alloc_ctor(0, 2, 0); } else { - x_97 = x_94; + x_96 = x_93; } -lean_ctor_set(x_97, 0, x_96); -lean_ctor_set(x_97, 1, x_93); -return x_97; +lean_ctor_set(x_96, 0, x_95); +lean_ctor_set(x_96, 1, x_92); +return x_96; } else { -lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; -lean_dec(x_88); -lean_dec(x_85); +lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; +lean_dec(x_87); lean_dec(x_84); lean_dec(x_83); -lean_free_object(x_40); -lean_dec(x_9); -x_98 = lean_ctor_get(x_91, 0); +lean_dec(x_82); +lean_free_object(x_39); +lean_dec(x_8); +x_97 = lean_ctor_get(x_90, 0); +lean_inc(x_97); +x_98 = lean_ctor_get(x_90, 1); lean_inc(x_98); -x_99 = lean_ctor_get(x_91, 1); -lean_inc(x_99); -if (lean_is_exclusive(x_91)) { - lean_ctor_release(x_91, 0); - lean_ctor_release(x_91, 1); - x_100 = x_91; +if (lean_is_exclusive(x_90)) { + lean_ctor_release(x_90, 0); + lean_ctor_release(x_90, 1); + x_99 = x_90; } else { - lean_dec_ref(x_91); - x_100 = lean_box(0); + lean_dec_ref(x_90); + x_99 = lean_box(0); } -if (lean_is_scalar(x_100)) { - x_101 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_99)) { + x_100 = lean_alloc_ctor(1, 2, 0); } else { - x_101 = x_100; + x_100 = x_99; } -lean_ctor_set(x_101, 0, x_98); -lean_ctor_set(x_101, 1, x_99); -return x_101; +lean_ctor_set(x_100, 0, x_97); +lean_ctor_set(x_100, 1, x_98); +return x_100; } } } else { -lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; uint8_t x_112; lean_object* x_113; -x_102 = lean_ctor_get(x_40, 0); +lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; uint8_t x_111; lean_object* x_112; +x_101 = lean_ctor_get(x_39, 0); +lean_inc(x_101); +lean_dec(x_39); +x_102 = lean_ctor_get(x_33, 1); lean_inc(x_102); -lean_dec(x_40); -x_103 = lean_ctor_get(x_34, 1); +lean_dec(x_33); +x_103 = lean_ctor_get(x_9, 0); lean_inc(x_103); -lean_dec(x_34); -x_104 = lean_ctor_get(x_10, 0); +x_104 = lean_ctor_get(x_9, 2); lean_inc(x_104); -x_105 = lean_ctor_get(x_10, 2); -lean_inc(x_105); -if (lean_is_exclusive(x_10)) { - lean_ctor_release(x_10, 0); - lean_ctor_release(x_10, 1); - lean_ctor_release(x_10, 2); - lean_ctor_release(x_10, 3); - x_106 = x_10; +if (lean_is_exclusive(x_9)) { + lean_ctor_release(x_9, 0); + lean_ctor_release(x_9, 1); + lean_ctor_release(x_9, 2); + lean_ctor_release(x_9, 3); + x_105 = x_9; } else { - lean_dec_ref(x_10); - x_106 = lean_box(0); + lean_dec_ref(x_9); + x_105 = lean_box(0); } -x_107 = lean_ctor_get(x_102, 0); +x_106 = lean_ctor_get(x_101, 0); +lean_inc(x_106); +x_107 = lean_ctor_get(x_101, 1); lean_inc(x_107); -x_108 = lean_ctor_get(x_102, 1); -lean_inc(x_108); -lean_inc(x_102); -x_109 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); -lean_closure_set(x_109, 0, x_102); -lean_closure_set(x_109, 1, x_3); -if (lean_is_exclusive(x_102)) { - lean_ctor_release(x_102, 0); - lean_ctor_release(x_102, 1); - x_110 = x_102; +lean_inc(x_101); +x_108 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process_parseHeader___lambda__7), 4, 2); +lean_closure_set(x_108, 0, x_101); +lean_closure_set(x_108, 1, x_3); +if (lean_is_exclusive(x_101)) { + lean_ctor_release(x_101, 0); + lean_ctor_release(x_101, 1); + x_109 = x_101; } else { - lean_dec_ref(x_102); - x_110 = lean_box(0); -} -x_111 = lean_ctor_get(x_108, 0); -lean_inc(x_111); -x_112 = 1; -x_113 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_108, x_109, x_111, x_112, x_103); -if (lean_obj_tag(x_113) == 0) -{ -lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; -x_114 = lean_ctor_get(x_113, 0); + lean_dec_ref(x_101); + x_109 = lean_box(0); +} +x_110 = lean_ctor_get(x_107, 0); +lean_inc(x_110); +x_111 = 1; +x_112 = l_Lean_Language_SnapshotTask_bindIO___rarg(x_107, x_108, x_110, x_111, x_102); +if (lean_obj_tag(x_112) == 0) +{ +lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; +x_113 = lean_ctor_get(x_112, 0); +lean_inc(x_113); +x_114 = lean_ctor_get(x_112, 1); lean_inc(x_114); -x_115 = lean_ctor_get(x_113, 1); -lean_inc(x_115); -if (lean_is_exclusive(x_113)) { - lean_ctor_release(x_113, 0); - lean_ctor_release(x_113, 1); - x_116 = x_113; +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_113); - x_116 = lean_box(0); + lean_dec_ref(x_112); + x_115 = lean_box(0); } -if (lean_is_scalar(x_110)) { - x_117 = lean_alloc_ctor(0, 2, 0); +if (lean_is_scalar(x_109)) { + x_116 = lean_alloc_ctor(0, 2, 0); } else { - x_117 = x_110; -} -lean_ctor_set(x_117, 0, x_107); -lean_ctor_set(x_117, 1, x_114); -x_118 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_118, 0, x_117); -if (lean_is_scalar(x_106)) { - x_119 = lean_alloc_ctor(0, 4, 0); + x_116 = x_109; +} +lean_ctor_set(x_116, 0, x_106); +lean_ctor_set(x_116, 1, x_113); +x_117 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_117, 0, x_116); +if (lean_is_scalar(x_105)) { + x_118 = lean_alloc_ctor(0, 4, 0); } else { - x_119 = x_106; -} -lean_ctor_set(x_119, 0, x_104); -lean_ctor_set(x_119, 1, x_9); -lean_ctor_set(x_119, 2, x_105); -lean_ctor_set(x_119, 3, x_118); -if (lean_is_scalar(x_116)) { - x_120 = lean_alloc_ctor(0, 2, 0); + x_118 = x_105; +} +lean_ctor_set(x_118, 0, x_103); +lean_ctor_set(x_118, 1, x_8); +lean_ctor_set(x_118, 2, x_104); +lean_ctor_set(x_118, 3, x_117); +if (lean_is_scalar(x_115)) { + x_119 = lean_alloc_ctor(0, 2, 0); } else { - x_120 = x_116; + x_119 = x_115; } -lean_ctor_set(x_120, 0, x_119); -lean_ctor_set(x_120, 1, x_115); -return x_120; +lean_ctor_set(x_119, 0, x_118); +lean_ctor_set(x_119, 1, x_114); +return x_119; } else { -lean_object* x_121; lean_object* x_122; lean_object* x_123; lean_object* x_124; -lean_dec(x_110); -lean_dec(x_107); +lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +lean_dec(x_109); lean_dec(x_106); lean_dec(x_105); lean_dec(x_104); -lean_dec(x_9); -x_121 = lean_ctor_get(x_113, 0); +lean_dec(x_103); +lean_dec(x_8); +x_120 = lean_ctor_get(x_112, 0); +lean_inc(x_120); +x_121 = lean_ctor_get(x_112, 1); lean_inc(x_121); -x_122 = lean_ctor_get(x_113, 1); -lean_inc(x_122); -if (lean_is_exclusive(x_113)) { - lean_ctor_release(x_113, 0); - lean_ctor_release(x_113, 1); - x_123 = x_113; +if (lean_is_exclusive(x_112)) { + lean_ctor_release(x_112, 0); + lean_ctor_release(x_112, 1); + x_122 = x_112; } else { - lean_dec_ref(x_113); - x_123 = lean_box(0); + lean_dec_ref(x_112); + x_122 = lean_box(0); } -if (lean_is_scalar(x_123)) { - x_124 = lean_alloc_ctor(1, 2, 0); +if (lean_is_scalar(x_122)) { + x_123 = lean_alloc_ctor(1, 2, 0); } else { - x_124 = x_123; + x_123 = x_122; } -lean_ctor_set(x_124, 0, x_121); -lean_ctor_set(x_124, 1, x_122); -return x_124; +lean_ctor_set(x_123, 0, x_120); +lean_ctor_set(x_123, 1, x_121); +return x_123; } } } @@ -7785,29 +7744,29 @@ return x_124; } else { -uint8_t x_125; -lean_dec(x_10); +uint8_t x_124; lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_125 = !lean_is_exclusive(x_24); -if (x_125 == 0) +x_124 = !lean_is_exclusive(x_23); +if (x_124 == 0) { -return x_24; +return x_23; } else { -lean_object* x_126; lean_object* x_127; lean_object* x_128; -x_126 = lean_ctor_get(x_24, 0); -x_127 = lean_ctor_get(x_24, 1); -lean_inc(x_127); +lean_object* x_125; lean_object* x_126; lean_object* x_127; +x_125 = lean_ctor_get(x_23, 0); +x_126 = lean_ctor_get(x_23, 1); lean_inc(x_126); -lean_dec(x_24); -x_128 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_128, 0, x_126); -lean_ctor_set(x_128, 1, x_127); -return x_128; +lean_inc(x_125); +lean_dec(x_23); +x_127 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_127, 0, x_125); +lean_ctor_set(x_127, 1, x_126); +return x_127; } } } @@ -7815,29 +7774,29 @@ return x_128; } else { -uint8_t x_129; -lean_dec(x_10); +uint8_t x_128; lean_dec(x_9); +lean_dec(x_8); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_129 = !lean_is_exclusive(x_12); -if (x_129 == 0) +x_128 = !lean_is_exclusive(x_11); +if (x_128 == 0) { -return x_12; +return x_11; } else { -lean_object* x_130; lean_object* x_131; lean_object* x_132; -x_130 = lean_ctor_get(x_12, 0); -x_131 = lean_ctor_get(x_12, 1); -lean_inc(x_131); +lean_object* x_129; lean_object* x_130; lean_object* x_131; +x_129 = lean_ctor_get(x_11, 0); +x_130 = lean_ctor_get(x_11, 1); lean_inc(x_130); -lean_dec(x_12); -x_132 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_132, 0, x_130); -lean_ctor_set(x_132, 1, x_131); -return x_132; +lean_inc(x_129); +lean_dec(x_11); +x_131 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_131, 0, x_129); +lean_ctor_set(x_131, 1, x_130); +return x_131; } } } @@ -7894,7 +7853,7 @@ return x_7; } else { -lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; x_8 = lean_ctor_get(x_2, 0); lean_inc(x_8); x_9 = lean_ctor_get(x_8, 1); @@ -7903,21 +7862,18 @@ lean_dec(x_8); x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); lean_dec(x_9); -x_11 = lean_ctor_get(x_3, 1); +x_11 = lean_ctor_get(x_3, 0); lean_inc(x_11); -x_12 = lean_ctor_get(x_11, 0); -lean_inc(x_12); +x_12 = l_String_firstDiffPos(x_10, x_11); lean_dec(x_11); -x_13 = l_String_firstDiffPos(x_10, x_12); -lean_dec(x_12); lean_dec(x_10); -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_3); -lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_Language_Lean_process_parseHeader(x_1, x_2, x_15, x_4); -return x_16; +x_13 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_13, 0, x_12); +x_14 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_14, 0, x_3); +lean_ctor_set(x_14, 1, x_13); +x_15 = l_Lean_Language_Lean_process_parseHeader(x_1, x_2, x_14, x_4); +return x_15; } } } @@ -8087,6 +8043,7 @@ l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__1 = _init_l_ lean_mark_persistent(l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__1); l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__2 = _init_l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__2(); lean_mark_persistent(l_Lean_Language_Lean_HeaderParsedSnapshot_processedResult___closed__2); +l_Lean_Language_Lean_SetupImportsResult_trustLevel___default = _init_l_Lean_Language_Lean_SetupImportsResult_trustLevel___default(); l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1 = _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1(); lean_mark_persistent(l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__1); l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2 = _init_l_IO_FS_withIsolatedStreams___at_Lean_Language_Lean_process_doElab___spec__1___closed__2(); @@ -8149,32 +8106,30 @@ l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___clos lean_mark_persistent(l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__2); l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3 = _init_l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3(); lean_mark_persistent(l_List_mapTR_loop___at_Lean_Language_Lean_process_processHeader___spec__1___closed__3); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__1); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__2); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__3); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__4); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__5); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__6); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__7); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__8); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__9); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__10); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__11); -l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12 = _init_l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__4___closed__12); -l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1(); -lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__5___closed__1); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__1); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__2); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__3); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__4); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__5); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__6); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__7); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__8); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__9); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__10); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__11); +l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12 = _init_l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12(); +lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___lambda__3___closed__12); l_Lean_Language_Lean_process_processHeader___closed__1 = _init_l_Lean_Language_Lean_process_processHeader___closed__1(); lean_mark_persistent(l_Lean_Language_Lean_process_processHeader___closed__1); l_Lean_Language_Lean_process_parseHeader___lambda__4___closed__1 = _init_l_Lean_Language_Lean_process_parseHeader___lambda__4___closed__1(); diff --git a/stage0/stdlib/Lean/Linter/MissingDocs.c b/stage0/stdlib/Lean/Linter/MissingDocs.c index c1fc7dfb083f..ecaaaf29dc83 100644 --- a/stage0/stdlib/Lean/Linter/MissingDocs.c +++ b/stage0/stdlib/Lean/Linter/MissingDocs.c @@ -143,6 +143,7 @@ lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkClassAbbrev___closed__1; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__6; LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkNotation(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkMacro___closed__2; size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_checkElab___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -341,7 +342,6 @@ lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_handleMutual___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkElab___closed__1; static lean_object* l_Lean_Linter_logLint___at_Lean_Linter_MissingDocs_lint___spec__1___closed__1; -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkRegisterOption___closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_MissingDocs_checkDecl___spec__15___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___regBuiltin_Lean_Linter_MissingDocs_checkSyntaxAbbrev___closed__4; @@ -2696,7 +2696,7 @@ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_Mis { uint8_t x_8; uint8_t x_9; x_8 = 0; -x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_8); +x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_8); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; @@ -3178,7 +3178,7 @@ LEAN_EXPORT lean_object* l_Lean_Linter_MissingDocs_initFn____x40_Lean_Linter_Mis { uint8_t x_8; uint8_t x_9; x_8 = 0; -x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_8); +x_9 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_8); if (x_9 == 0) { lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; uint8_t x_16; diff --git a/stage0/stdlib/Lean/Linter/UnusedVariables.c b/stage0/stdlib/Lean/Linter/UnusedVariables.c index 253822bd79f8..7bff57c2987b 100644 --- a/stage0/stdlib/Lean/Linter/UnusedVariables.c +++ b/stage0/stdlib/Lean/Linter/UnusedVariables.c @@ -201,6 +201,7 @@ LEAN_EXPORT lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables uint64_t l___private_Lean_Syntax_0__String_hashRange____x40_Lean_Syntax___hyg_168_(lean_object*); static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1273____spec__3___closed__4; LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Linter_UnusedVariables_collectReferences___spec__19(lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_ForEachExpr_visit___at_Lean_Linter_UnusedVariables_visitAssignments_visitExpr___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l_Lean_Linter_getLinterUnusedVariablesPatternVars(lean_object*); @@ -476,7 +477,6 @@ LEAN_EXPORT uint8_t l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hy static lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1273____spec__3___closed__6; static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_2420____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Linter_UnusedVariables_unusedVariables___elambda__1___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_593____closed__4; static lean_object* l_Array_qsort_sort___at_Lean_Linter_UnusedVariables_unusedVariables___elambda__1___spec__19___closed__1; LEAN_EXPORT lean_object* l_Array_anyMUnsafe_any___at_Lean_Linter_initFn____x40_Lean_Linter_UnusedVariables___hyg_1273____spec__3___boxed(lean_object*, lean_object*, lean_object*); @@ -2132,7 +2132,7 @@ x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); x_10 = 0; -x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_10); +x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; @@ -2439,7 +2439,7 @@ x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); x_10 = 0; -x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_10); +x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; diff --git a/stage0/stdlib/Lean/Meta/Closure.c b/stage0/stdlib/Lean/Meta/Closure.c index 579c3001d432..f44c7ce8bb60 100644 --- a/stage0/stdlib/Lean/Meta/Closure.c +++ b/stage0/stdlib/Lean/Meta/Closure.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Closure -// Imports: Lean.MetavarContext Lean.Environment Lean.Util.FoldConsts Lean.Meta.Basic Lean.Meta.Check +// Imports: Lean.MetavarContext Lean.Environment Lean.AddDecl Lean.Util.FoldConsts Lean.Meta.Basic Lean.Meta.Check #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -12168,6 +12168,7 @@ return x_10; } lean_object* initialize_Lean_MetavarContext(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Environment(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_FoldConsts(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Check(uint8_t builtin, lean_object*); @@ -12182,6 +12183,9 @@ lean_dec_ref(res); res = initialize_Lean_Environment(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Util_FoldConsts(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Meta/CoeAttr.c b/stage0/stdlib/Lean/Meta/CoeAttr.c index ee81949e391d..7362fc847bc8 100644 --- a/stage0/stdlib/Lean/Meta/CoeAttr.c +++ b/stage0/stdlib/Lean/Meta/CoeAttr.c @@ -97,6 +97,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_CoeAttr_0__Lean_Meta_reprCoeFnInf lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_registerCoercion___lambda__1___closed__15; lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_CoeAttr___hyg_436____closed__5; static lean_object* l___private_Lean_Meta_CoeAttr_0__Lean_Meta_reprCoeFnType____x40_Lean_Meta_CoeAttr___hyg_16____closed__14; static lean_object* l_Lean_Meta_instToExprCoeFnInfo___lambda__1___closed__3; @@ -198,7 +199,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_instDecidableEqCoeFnType___boxed(lean_objec static lean_object* l___private_Lean_Meta_CoeAttr_0__Lean_Meta_reprCoeFnType____x40_Lean_Meta_CoeAttr___hyg_16____closed__11; static lean_object* l_Lean_Meta_instToExprCoeFnType___lambda__1___closed__11; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_CoeAttr___hyg_763____closed__7; -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Meta_instReprCoeFnInfo; static lean_object* l___private_Lean_Meta_CoeAttr_0__Lean_Meta_reprCoeFnType____x40_Lean_Meta_CoeAttr___hyg_16____closed__15; lean_object* l_Lean_addMessageContextFull___at_Lean_Meta_instAddMessageContextMetaM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -2588,7 +2588,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_CoeAttr___hyg_763__ { uint8_t x_7; uint8_t x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; x_7 = 0; -x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_7); +x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_7); x_9 = l_Lean_Meta_initFn____x40_Lean_Meta_CoeAttr___hyg_763____lambda__2___closed__13; x_10 = lean_st_mk_ref(x_9, x_6); if (x_8 == 0) diff --git a/stage0/stdlib/Lean/Meta/Constructions.c b/stage0/stdlib/Lean/Meta/Constructions.c index 41e3c5c810e3..0e7767bbaab2 100644 --- a/stage0/stdlib/Lean/Meta/Constructions.c +++ b/stage0/stdlib/Lean/Meta/Constructions.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Constructions -// Imports: Lean.AuxRecursor Lean.Meta.AppBuilder +// Imports: Lean.AuxRecursor Lean.AddDecl Lean.Meta.AppBuilder #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -1357,7 +1357,7 @@ static lean_object* _init_l_Lean_mkNoConfusionEnum_mkToCtorIdx___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_mkNoConfusionEnum_mkToCtorIdx___closed__1; x_2 = l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__2; -x_3 = lean_unsigned_to_nat(47u); +x_3 = lean_unsigned_to_nat(48u); x_4 = lean_unsigned_to_nat(63u); x_5 = l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -1920,7 +1920,7 @@ static lean_object* _init_l_Lean_mkNoConfusionEnum_mkNoConfusionType___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_mkNoConfusionEnum_mkToCtorIdx___closed__1; x_2 = l_Lean_mkNoConfusionEnum_mkNoConfusionType___closed__1; -x_3 = lean_unsigned_to_nat(71u); +x_3 = lean_unsigned_to_nat(72u); x_4 = lean_unsigned_to_nat(63u); x_5 = l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -2720,7 +2720,7 @@ static lean_object* _init_l_Lean_mkNoConfusionEnum_mkNoConfusion___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_mkNoConfusionEnum_mkToCtorIdx___closed__1; x_2 = l_Lean_mkNoConfusionEnum_mkNoConfusion___closed__1; -x_3 = lean_unsigned_to_nat(94u); +x_3 = lean_unsigned_to_nat(95u); x_4 = lean_unsigned_to_nat(63u); x_5 = l_Lean_mkNoConfusionEnum_mkToCtorIdx___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -3884,6 +3884,7 @@ return x_7; } } lean_object* initialize_Lean_AuxRecursor(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(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_Meta_Constructions(uint8_t builtin, lean_object* w) { @@ -3893,6 +3894,9 @@ _G_initialized = true; res = initialize_Lean_AuxRecursor(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(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()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Meta/Eqns.c b/stage0/stdlib/Lean/Meta/Eqns.c index 1cb2a8cf63cb..787c9f964219 100644 --- a/stage0/stdlib/Lean/Meta/Eqns.c +++ b/stage0/stdlib/Lean/Meta/Eqns.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Eqns -// Imports: Lean.ReservedNameAction Lean.Meta.Basic Lean.Meta.AppBuilder Lean.Meta.Match.MatcherInfo +// Imports: Lean.ReservedNameAction Lean.AddDecl Lean.Meta.Basic Lean.Meta.AppBuilder Lean.Meta.Match.MatcherInfo #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -6712,6 +6712,7 @@ return x_7; } } lean_object* initialize_Lean_ReservedNameAction(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_AppBuilder(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Match_MatcherInfo(uint8_t builtin, lean_object*); @@ -6723,6 +6724,9 @@ _G_initialized = true; res = initialize_Lean_ReservedNameAction(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(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); diff --git a/stage0/stdlib/Lean/Meta/Eval.c b/stage0/stdlib/Lean/Meta/Eval.c index fa2e2dd0cec2..a4b7feea57eb 100644 --- a/stage0/stdlib/Lean/Meta/Eval.c +++ b/stage0/stdlib/Lean/Meta/Eval.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Eval -// Imports: Lean.Meta.Check +// Imports: Lean.AddDecl Lean.Meta.Check #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -967,12 +967,16 @@ x_10 = l_Lean_Meta_evalExpr___rarg(x_1, x_2, x_9, x_4, x_5, x_6, x_7, x_8); return x_10; } } +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Check(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Eval(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_AddDecl(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Meta_Check(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Meta/Instances.c b/stage0/stdlib/Lean/Meta/Instances.c index 652b677f4c6f..cc2673cea2eb 100644 --- a/stage0/stdlib/Lean/Meta/Instances.c +++ b/stage0/stdlib/Lean/Meta/Instances.c @@ -162,6 +162,7 @@ static lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrd lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Instances___hyg_2534____closed__15; LEAN_EXPORT uint8_t l_Lean_Meta_isInstanceCore(lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Meta_Instances_erase___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__8(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*); @@ -373,7 +374,6 @@ lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addInstanceEntry___spec__6(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn___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_Instances___hyg_2534____closed__20; -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); static lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___lambda__1___closed__8; LEAN_EXPORT lean_object* l___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder_assignMVarsIn(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Meta_Instances_0__Lean_Meta_computeSynthOrder___spec__8___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -16540,7 +16540,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = 0; -x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_12); +x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_12); if (x_13 == 0) { lean_object* x_14; lean_object* x_15; uint8_t x_16; diff --git a/stage0/stdlib/Lean/Meta/LazyDiscrTree.c b/stage0/stdlib/Lean/Meta/LazyDiscrTree.c index 94ae1e849771..c626ac6a41c7 100644 --- a/stage0/stdlib/Lean/Meta/LazyDiscrTree.c +++ b/stage0/stdlib/Lean/Meta/LazyDiscrTree.c @@ -18,6 +18,7 @@ static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTr static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_reprKey____x40_Lean_Meta_LazyDiscrTree___hyg_367____closed__7; static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_isNumeral___closed__4; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntry___spec__3___boxed(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___rarg___lambda__1(lean_object*, lean_object*, lean_object*); @@ -69,7 +70,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_findImportMatches(lean_object LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_InitEntry_mkSubEntry___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_logAt___at_Lean_Meta_LazyDiscrTree_findImportMatches___spec__7(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_pushArgs(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_log___at_Lean_addDecl___spec__5(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_reprKey____x40_Lean_Meta_LazyDiscrTree___hyg_367____closed__29; LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_setTrie___rarg(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*); @@ -82,6 +82,7 @@ lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_PreDiscrTree_push(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntries___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 l_Lean_Expr_isRawNatLit(lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_findMatchesExt___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_LazyDiscrTree_tries___default___closed__1; lean_object* lean_io_as_task(lean_object*, lean_object*, lean_object*); @@ -129,6 +130,7 @@ lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createModuleTreeRef___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Expr_instBEq; +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_newTrie___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntry___spec__11___rarg___boxed(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); @@ -150,7 +152,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_MatchResult_appendResults___r LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_PreDiscrTree_append___spec__2(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntry___spec__6___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_ignoreArg___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_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_addConstImportData___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*); uint8_t l_Lean_ConstantInfo_isUnsafe(lean_object*); @@ -257,6 +258,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDi static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_isNatOffset___closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree_go___at_Lean_Meta_LazyDiscrTree_findImportMatches___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*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_PreDiscrTree_toLazy___spec__1___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__4(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_getKeyArgs(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_ImportData_new(lean_object*); static lean_object* l_Lean_Meta_LazyDiscrTree_blacklistInsertion___closed__3; @@ -279,6 +281,7 @@ static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTr static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_getKeyArgs___lambda__3___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_getMatchCore(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___at_Lean_Meta_LazyDiscrTree_findImportMatches___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_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createImportedDiscrTree___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_PreDiscrTree_instAppendPreDiscrTree(lean_object*); @@ -297,6 +300,7 @@ lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_PreDiscrTree_toLazy___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_LazyDiscrTree_PreDiscrTree_append___spec__1(lean_object*); LEAN_EXPORT uint64_t l_Lean_Meta_LazyDiscrTree_Key_hash(lean_object*); +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntry___spec__7(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_LazyDiscrTree_isInternalDetail(lean_object*); uint8_t l_Lean_Expr_hasLooseBVars(lean_object*); @@ -309,6 +313,7 @@ static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTr LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_MatchResult_size___spec__2(lean_object*); static lean_object* l_Lean_Meta_LazyDiscrTree_Cache_empty___closed__5; static lean_object* l_Lean_Meta_LazyDiscrTree_PreDiscrTree_append___rarg___closed__1; +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_getKeyArgs___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_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_runMatch___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_findModuleMatches___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -365,7 +370,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_getChildNgen(lean_object*); static lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_elimLooseBVarsByBeta___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_findImportMatches___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_createImportedEnvironmentSeq_go___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_MatchClone_isNumeral(lean_object*); lean_object* l_Lean_MVarId_isReadOnlyOrSyntheticOpaque(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_LazyDiscrTree_0__Lean_Meta_LazyDiscrTree_evalLazyEntry___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -20159,6 +20163,909 @@ x_2 = lean_alloc_closure((void*)(l_Lean_Meta_LazyDiscrTree_getChildNgen___at_Lea return x_2; } } +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__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) { +_start: +{ +lean_object* x_7; uint8_t x_265; uint8_t x_266; +x_265 = 2; +x_266 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_265); +if (x_266 == 0) +{ +lean_object* x_267; +x_267 = lean_box(0); +x_7 = x_267; +goto block_264; +} +else +{ +lean_object* x_268; uint8_t x_269; +lean_inc(x_2); +x_268 = l_Lean_MessageData_hasSyntheticSorry(x_2); +x_269 = lean_unbox(x_268); +lean_dec(x_268); +if (x_269 == 0) +{ +lean_object* x_270; +x_270 = lean_box(0); +x_7 = x_270; +goto block_264; +} +else +{ +lean_object* x_271; lean_object* x_272; +lean_dec(x_2); +x_271 = lean_box(0); +x_272 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_6); +return x_272; +} +} +block_264: +{ +uint8_t x_8; lean_object* x_258; uint8_t x_259; uint8_t x_260; +lean_dec(x_7); +x_258 = lean_ctor_get(x_4, 2); +x_259 = 1; +x_260 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_259); +if (x_260 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +lean_object* x_261; uint8_t x_262; +x_261 = l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_findImportMatches___spec__7___closed__1; +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); +if (x_262 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +uint8_t x_263; +x_263 = 2; +x_8 = x_263; +goto block_257; +} +} +block_257: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_ctor_get(x_4, 5); +x_12 = lean_ctor_get(x_4, 6); +x_13 = lean_ctor_get(x_4, 7); +x_14 = l_Lean_replaceRef(x_1, x_11); +x_15 = 0; +x_16 = l_Lean_Syntax_getPos_x3f(x_14, x_15); +x_17 = l_Lean_Syntax_getTailPos_x3f(x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_18 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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_unsigned_to_nat(0u); +x_22 = l_Lean_FileMap_toPosition(x_10, x_21); +lean_inc(x_22); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_inc(x_13); +lean_inc(x_12); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_13); +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +x_26 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_27 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 3, x_26); +lean_ctor_set(x_27, 4, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_8); +x_28 = lean_st_ref_take(x_5, x_20); +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; uint8_t x_35; +x_32 = lean_ctor_get(x_29, 5); +x_33 = l_Lean_PersistentArray_push___rarg(x_32, x_27); +lean_ctor_set(x_29, 5, x_33); +x_34 = lean_st_ref_set(x_5, x_29, x_30); +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_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_29, 0); +x_42 = lean_ctor_get(x_29, 1); +x_43 = lean_ctor_get(x_29, 2); +x_44 = lean_ctor_get(x_29, 3); +x_45 = lean_ctor_get(x_29, 4); +x_46 = lean_ctor_get(x_29, 5); +x_47 = lean_ctor_get(x_29, 6); +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_inc(x_41); +lean_dec(x_29); +x_48 = l_Lean_PersistentArray_push___rarg(x_46, x_27); +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_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_48); +lean_ctor_set(x_49, 6, x_47); +x_50 = lean_st_ref_set(x_5, x_49, x_30); +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 +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_17); +if (x_55 == 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; lean_object* x_69; uint8_t x_70; +x_56 = lean_ctor_get(x_17, 0); +x_57 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l_Lean_FileMap_toPosition(x_10, x_60); +x_62 = l_Lean_FileMap_toPosition(x_10, x_56); +lean_dec(x_56); +lean_ctor_set(x_17, 0, x_62); +lean_inc(x_13); +lean_inc(x_12); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_12); +lean_ctor_set(x_63, 1, x_13); +x_64 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +x_65 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_66 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_66, 0, x_9); +lean_ctor_set(x_66, 1, x_61); +lean_ctor_set(x_66, 2, x_17); +lean_ctor_set(x_66, 3, x_65); +lean_ctor_set(x_66, 4, x_64); +lean_ctor_set_uint8(x_66, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_66, sizeof(void*)*5 + 1, x_8); +x_67 = lean_st_ref_take(x_5, x_59); +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_is_exclusive(x_68); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_68, 5); +x_72 = l_Lean_PersistentArray_push___rarg(x_71, x_66); +lean_ctor_set(x_68, 5, x_72); +x_73 = lean_st_ref_set(x_5, x_68, x_69); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +x_76 = lean_box(0); +lean_ctor_set(x_73, 0, x_76); +return x_73; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_80 = lean_ctor_get(x_68, 0); +x_81 = lean_ctor_get(x_68, 1); +x_82 = lean_ctor_get(x_68, 2); +x_83 = lean_ctor_get(x_68, 3); +x_84 = lean_ctor_get(x_68, 4); +x_85 = lean_ctor_get(x_68, 5); +x_86 = lean_ctor_get(x_68, 6); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_68); +x_87 = l_Lean_PersistentArray_push___rarg(x_85, x_66); +x_88 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 1, x_81); +lean_ctor_set(x_88, 2, x_82); +lean_ctor_set(x_88, 3, x_83); +lean_ctor_set(x_88, 4, x_84); +lean_ctor_set(x_88, 5, x_87); +lean_ctor_set(x_88, 6, x_86); +x_89 = lean_st_ref_set(x_5, x_88, x_69); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_91 = x_89; +} else { + lean_dec_ref(x_89); + x_91 = lean_box(0); +} +x_92 = lean_box(0); +if (lean_is_scalar(x_91)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_91; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +return x_93; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_94 = lean_ctor_get(x_17, 0); +lean_inc(x_94); +lean_dec(x_17); +x_95 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l_Lean_FileMap_toPosition(x_10, x_98); +x_100 = l_Lean_FileMap_toPosition(x_10, x_94); +lean_dec(x_94); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +lean_inc(x_13); +lean_inc(x_12); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_12); +lean_ctor_set(x_102, 1, x_13); +x_103 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_96); +x_104 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_105 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_105, 0, x_9); +lean_ctor_set(x_105, 1, x_99); +lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 3, x_104); +lean_ctor_set(x_105, 4, x_103); +lean_ctor_set_uint8(x_105, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_105, sizeof(void*)*5 + 1, x_8); +x_106 = lean_st_ref_take(x_5, x_97); +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_ctor_get(x_107, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +x_112 = lean_ctor_get(x_107, 3); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 4); +lean_inc(x_113); +x_114 = lean_ctor_get(x_107, 5); +lean_inc(x_114); +x_115 = lean_ctor_get(x_107, 6); +lean_inc(x_115); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + lean_ctor_release(x_107, 3); + lean_ctor_release(x_107, 4); + lean_ctor_release(x_107, 5); + lean_ctor_release(x_107, 6); + x_116 = x_107; +} else { + lean_dec_ref(x_107); + x_116 = lean_box(0); +} +x_117 = l_Lean_PersistentArray_push___rarg(x_114, x_105); +if (lean_is_scalar(x_116)) { + x_118 = lean_alloc_ctor(0, 7, 0); +} else { + x_118 = x_116; +} +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_110); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set(x_118, 3, x_112); +lean_ctor_set(x_118, 4, x_113); +lean_ctor_set(x_118, 5, x_117); +lean_ctor_set(x_118, 6, x_115); +x_119 = lean_st_ref_set(x_5, x_118, x_108); +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 +{ +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_124; +x_124 = !lean_is_exclusive(x_16); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_125 = lean_ctor_get(x_16, 0); +x_126 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_FileMap_toPosition(x_10, x_125); +lean_dec(x_125); +lean_inc(x_129); +lean_ctor_set(x_16, 0, x_129); +lean_inc(x_13); +lean_inc(x_12); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_12); +lean_ctor_set(x_130, 1, x_13); +x_131 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_127); +x_132 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_133 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_133, 0, x_9); +lean_ctor_set(x_133, 1, x_129); +lean_ctor_set(x_133, 2, x_16); +lean_ctor_set(x_133, 3, x_132); +lean_ctor_set(x_133, 4, x_131); +lean_ctor_set_uint8(x_133, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_133, sizeof(void*)*5 + 1, x_8); +x_134 = lean_st_ref_take(x_5, x_128); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = !lean_is_exclusive(x_135); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_138 = lean_ctor_get(x_135, 5); +x_139 = l_Lean_PersistentArray_push___rarg(x_138, x_133); +lean_ctor_set(x_135, 5, x_139); +x_140 = lean_st_ref_set(x_5, x_135, x_136); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_140, 0); +lean_dec(x_142); +x_143 = lean_box(0); +lean_ctor_set(x_140, 0, x_143); +return x_140; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_140, 1); +lean_inc(x_144); +lean_dec(x_140); +x_145 = lean_box(0); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_147 = lean_ctor_get(x_135, 0); +x_148 = lean_ctor_get(x_135, 1); +x_149 = lean_ctor_get(x_135, 2); +x_150 = lean_ctor_get(x_135, 3); +x_151 = lean_ctor_get(x_135, 4); +x_152 = lean_ctor_get(x_135, 5); +x_153 = lean_ctor_get(x_135, 6); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_135); +x_154 = l_Lean_PersistentArray_push___rarg(x_152, x_133); +x_155 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_155, 0, x_147); +lean_ctor_set(x_155, 1, x_148); +lean_ctor_set(x_155, 2, x_149); +lean_ctor_set(x_155, 3, x_150); +lean_ctor_set(x_155, 4, x_151); +lean_ctor_set(x_155, 5, x_154); +lean_ctor_set(x_155, 6, x_153); +x_156 = lean_st_ref_set(x_5, x_155, x_136); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_158 = x_156; +} else { + lean_dec_ref(x_156); + x_158 = lean_box(0); +} +x_159 = lean_box(0); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_158; +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_161 = lean_ctor_get(x_16, 0); +lean_inc(x_161); +lean_dec(x_16); +x_162 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_FileMap_toPosition(x_10, x_161); +lean_dec(x_161); +lean_inc(x_165); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +lean_inc(x_13); +lean_inc(x_12); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_12); +lean_ctor_set(x_167, 1, x_13); +x_168 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_163); +x_169 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_170 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_170, 0, x_9); +lean_ctor_set(x_170, 1, x_165); +lean_ctor_set(x_170, 2, x_166); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_168); +lean_ctor_set_uint8(x_170, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_170, sizeof(void*)*5 + 1, x_8); +x_171 = lean_st_ref_take(x_5, x_164); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_172, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_172, 2); +lean_inc(x_176); +x_177 = lean_ctor_get(x_172, 3); +lean_inc(x_177); +x_178 = lean_ctor_get(x_172, 4); +lean_inc(x_178); +x_179 = lean_ctor_get(x_172, 5); +lean_inc(x_179); +x_180 = lean_ctor_get(x_172, 6); +lean_inc(x_180); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + lean_ctor_release(x_172, 2); + lean_ctor_release(x_172, 3); + lean_ctor_release(x_172, 4); + lean_ctor_release(x_172, 5); + lean_ctor_release(x_172, 6); + x_181 = x_172; +} else { + lean_dec_ref(x_172); + x_181 = lean_box(0); +} +x_182 = l_Lean_PersistentArray_push___rarg(x_179, x_170); +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 7, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_174); +lean_ctor_set(x_183, 1, x_175); +lean_ctor_set(x_183, 2, x_176); +lean_ctor_set(x_183, 3, x_177); +lean_ctor_set(x_183, 4, x_178); +lean_ctor_set(x_183, 5, x_182); +lean_ctor_set(x_183, 6, x_180); +x_184 = lean_st_ref_set(x_5, x_183, x_173); +x_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = lean_box(0); +if (lean_is_scalar(x_186)) { + x_188 = lean_alloc_ctor(0, 2, 0); +} else { + x_188 = x_186; +} +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_185); +return x_188; +} +} +else +{ +lean_object* x_189; uint8_t x_190; +x_189 = lean_ctor_get(x_16, 0); +lean_inc(x_189); +lean_dec(x_16); +x_190 = !lean_is_exclusive(x_17); +if (x_190 == 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; 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; uint8_t x_204; +x_191 = lean_ctor_get(x_17, 0); +x_192 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_196 = l_Lean_FileMap_toPosition(x_10, x_191); +lean_dec(x_191); +lean_ctor_set(x_17, 0, x_196); +lean_inc(x_13); +lean_inc(x_12); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_12); +lean_ctor_set(x_197, 1, x_13); +x_198 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_193); +x_199 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_200 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_200, 0, x_9); +lean_ctor_set(x_200, 1, x_195); +lean_ctor_set(x_200, 2, x_17); +lean_ctor_set(x_200, 3, x_199); +lean_ctor_set(x_200, 4, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_200, sizeof(void*)*5 + 1, x_8); +x_201 = lean_st_ref_take(x_5, x_194); +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_is_exclusive(x_202); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; +x_205 = lean_ctor_get(x_202, 5); +x_206 = l_Lean_PersistentArray_push___rarg(x_205, x_200); +lean_ctor_set(x_202, 5, x_206); +x_207 = lean_st_ref_set(x_5, x_202, x_203); +x_208 = !lean_is_exclusive(x_207); +if (x_208 == 0) +{ +lean_object* x_209; lean_object* x_210; +x_209 = lean_ctor_get(x_207, 0); +lean_dec(x_209); +x_210 = lean_box(0); +lean_ctor_set(x_207, 0, x_210); +return x_207; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +lean_dec(x_207); +x_212 = lean_box(0); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_211); +return x_213; +} +} +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; +x_214 = lean_ctor_get(x_202, 0); +x_215 = lean_ctor_get(x_202, 1); +x_216 = lean_ctor_get(x_202, 2); +x_217 = lean_ctor_get(x_202, 3); +x_218 = lean_ctor_get(x_202, 4); +x_219 = lean_ctor_get(x_202, 5); +x_220 = lean_ctor_get(x_202, 6); +lean_inc(x_220); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_202); +x_221 = l_Lean_PersistentArray_push___rarg(x_219, x_200); +x_222 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_222, 0, x_214); +lean_ctor_set(x_222, 1, x_215); +lean_ctor_set(x_222, 2, x_216); +lean_ctor_set(x_222, 3, x_217); +lean_ctor_set(x_222, 4, x_218); +lean_ctor_set(x_222, 5, x_221); +lean_ctor_set(x_222, 6, x_220); +x_223 = lean_st_ref_set(x_5, x_222, x_203); +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_223); + x_225 = lean_box(0); +} +x_226 = lean_box(0); +if (lean_is_scalar(x_225)) { + x_227 = lean_alloc_ctor(0, 2, 0); +} else { + x_227 = x_225; +} +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_224); +return x_227; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_228 = lean_ctor_get(x_17, 0); +lean_inc(x_228); +lean_dec(x_17); +x_229 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_233 = l_Lean_FileMap_toPosition(x_10, x_228); +lean_dec(x_228); +x_234 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_234, 0, x_233); +lean_inc(x_13); +lean_inc(x_12); +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_12); +lean_ctor_set(x_235, 1, x_13); +x_236 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_236, 0, x_235); +lean_ctor_set(x_236, 1, x_230); +x_237 = l_Lean_Meta_LazyDiscrTree_logImportFailure___rarg___closed__7; +lean_inc(x_9); +x_238 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_238, 0, x_9); +lean_ctor_set(x_238, 1, x_232); +lean_ctor_set(x_238, 2, x_234); +lean_ctor_set(x_238, 3, x_237); +lean_ctor_set(x_238, 4, x_236); +lean_ctor_set_uint8(x_238, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_238, sizeof(void*)*5 + 1, x_8); +x_239 = lean_st_ref_take(x_5, x_231); +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 = lean_ctor_get(x_240, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_240, 1); +lean_inc(x_243); +x_244 = lean_ctor_get(x_240, 2); +lean_inc(x_244); +x_245 = lean_ctor_get(x_240, 3); +lean_inc(x_245); +x_246 = lean_ctor_get(x_240, 4); +lean_inc(x_246); +x_247 = lean_ctor_get(x_240, 5); +lean_inc(x_247); +x_248 = lean_ctor_get(x_240, 6); +lean_inc(x_248); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + lean_ctor_release(x_240, 2); + lean_ctor_release(x_240, 3); + lean_ctor_release(x_240, 4); + lean_ctor_release(x_240, 5); + lean_ctor_release(x_240, 6); + x_249 = x_240; +} else { + lean_dec_ref(x_240); + x_249 = lean_box(0); +} +x_250 = l_Lean_PersistentArray_push___rarg(x_247, x_238); +if (lean_is_scalar(x_249)) { + x_251 = lean_alloc_ctor(0, 7, 0); +} else { + x_251 = x_249; +} +lean_ctor_set(x_251, 0, x_242); +lean_ctor_set(x_251, 1, x_243); +lean_ctor_set(x_251, 2, x_244); +lean_ctor_set(x_251, 3, x_245); +lean_ctor_set(x_251, 4, x_246); +lean_ctor_set(x_251, 5, x_250); +lean_ctor_set(x_251, 6, x_248); +x_252 = lean_st_ref_set(x_5, x_251, x_241); +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_254 = x_252; +} else { + lean_dec_ref(x_252); + x_254 = lean_box(0); +} +x_255 = lean_box(0); +if (lean_is_scalar(x_254)) { + x_256 = lean_alloc_ctor(0, 2, 0); +} else { + x_256 = x_254; +} +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_253); +return x_256; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 5); +lean_inc(x_6); +x_7 = l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__4(x_6, x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_6); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_Meta_LazyDiscrTree_logImportFailure___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -20196,11 +21103,11 @@ x_20 = lean_alloc_ctor(7, 2, 0); lean_ctor_set(x_20, 0, x_18); lean_ctor_set(x_20, 1, x_19); x_21 = 2; -x_22 = l_Lean_log___at_Lean_addDecl___spec__5(x_20, x_21, x_2, x_3, x_4); +x_22 = l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(x_20, x_21, x_2, x_3, x_4); return x_22; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___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_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5(lean_object* x_1, size_t x_2, size_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { uint8_t x_8; @@ -20317,7 +21224,7 @@ x_30 = 0; x_31 = lean_usize_of_nat(x_22); lean_dec(x_22); x_32 = lean_box(0); -x_33 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(x_20, x_30, x_31, x_32, x_2, x_3, x_21); +x_33 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5(x_20, x_30, x_31, x_32, x_2, x_3, x_21); lean_dec(x_20); x_34 = !lean_is_exclusive(x_33); if (x_34 == 0) @@ -20396,7 +21303,7 @@ x_54 = 0; x_55 = lean_usize_of_nat(x_44); lean_dec(x_44); x_56 = lean_box(0); -x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(x_42, x_54, x_55, x_56, x_2, x_3, x_43); +x_57 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5(x_42, x_54, x_55, x_56, x_2, x_3, x_43); lean_dec(x_42); x_58 = lean_ctor_get(x_57, 1); lean_inc(x_58); @@ -20449,7 +21356,30 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_logAt___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__4(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__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_6; lean_object* x_7; +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = l_Lean_log___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(x_1, x_6, x_3, x_4, x_5); +return x_7; +} +} +LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { size_t x_8; size_t x_9; lean_object* x_10; @@ -20457,7 +21387,7 @@ x_8 = lean_unbox_usize(x_2); lean_dec(x_2); x_9 = lean_unbox_usize(x_3); lean_dec(x_3); -x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__3(x_1, x_8, x_9, x_4, x_5, x_6, x_7); +x_10 = l_Array_foldlMUnsafe_fold___at_Lean_Meta_LazyDiscrTree_createModuleDiscrTree___spec__5(x_1, x_8, x_9, x_4, x_5, x_6, x_7); lean_dec(x_1); return x_10; } diff --git a/stage0/stdlib/Lean/Meta/SizeOf.c b/stage0/stdlib/Lean/Meta/SizeOf.c index d889327cc26b..0c15760cecda 100644 --- a/stage0/stdlib/Lean/Meta/SizeOf.c +++ b/stage0/stdlib/Lean/Meta/SizeOf.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.SizeOf -// Imports: Init.Data.List.BasicAux Lean.Meta.AppBuilder Lean.Meta.Instances +// Imports: Init.Data.List.BasicAux Lean.AddDecl Lean.Meta.AppBuilder Lean.Meta.Instances #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -4842,7 +4842,7 @@ static lean_object* _init_l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMino lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; lean_object* x_6; x_1 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__4; x_2 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__5; -x_3 = lean_unsigned_to_nat(101u); +x_3 = lean_unsigned_to_nat(102u); x_4 = lean_unsigned_to_nat(2u); x_5 = l___private_Lean_Meta_SizeOf_0__Lean_Meta_mkSizeOfMinors___rarg___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -16813,6 +16813,7 @@ return x_5; } } lean_object* initialize_Init_Data_List_BasicAux(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_AppBuilder(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Instances(uint8_t builtin, lean_object*); static bool _G_initialized = false; @@ -16823,6 +16824,9 @@ _G_initialized = true; res = initialize_Init_Data_List_BasicAux(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(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()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/Meta/SynthInstance.c b/stage0/stdlib/Lean/Meta/SynthInstance.c index 5b104f087b0e..63d514b8b75e 100644 --- a/stage0/stdlib/Lean/Meta/SynthInstance.c +++ b/stage0/stdlib/Lean/Meta/SynthInstance.c @@ -30,7 +30,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___lambda__1(lean_ob static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_withMCtx___at_Lean_Meta_SynthInstance_tryResolve___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams(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_9892____closed__7; LEAN_EXPORT lean_object* l_Lean_profileitM___at_Lean_Meta_synthInstance_x3f___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_Meta_SynthInstance_checkSystem___boxed(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__4; @@ -63,7 +62,6 @@ 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_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__8; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_generate___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_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; @@ -79,6 +77,7 @@ lean_object* l_Lean_profileitIOUnsafe___rarg(lean_object*, lean_object*, lean_ob static lean_object* l_Lean_Meta_SynthInstance_checkSystem___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_check_canceled(lean_object*); lean_object* lean_whnf(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -92,7 +91,6 @@ uint8_t lean_usize_dec_le(size_t, size_t); 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; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__10; 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; @@ -138,7 +136,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM; lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_SynthInstance_newSubgoal___spec__5(lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, 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*); @@ -161,7 +158,7 @@ LEAN_EXPORT lean_object* l_ReaderT_read___at_Lean_Meta_SynthInstance_addAnswer__ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_consume___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_float_decLt(double, double); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__4___closed__1; -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__6; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_SynthInstance_main___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -194,7 +191,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__5(lean_o LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_addAnswer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__6; 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; 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*); @@ -204,8 +200,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__1___b 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*); 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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__5; 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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__10; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__1; static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_82____closed__7; 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*); @@ -218,7 +214,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSu 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*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__6; 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; @@ -236,6 +231,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__3(lean_o lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__5; static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__1___closed__4; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__13; lean_object* l_Lean_MVarId_assign___at_Lean_Meta_getLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__2; size_t lean_ptr_addr(lean_object*); @@ -260,6 +256,7 @@ LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_SynthInstance_get static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__5; LEAN_EXPORT lean_object* l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__7; lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getEntry(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance_x3f___spec__7___closed__2; @@ -269,6 +266,7 @@ static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_82__ lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_processPostponed___spec__2___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessOutParam(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__12; lean_object* l_Lean_getOutParamPositions_x3f(lean_object*, lean_object*); uint8_t lean_expr_eqv(lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_main___lambda__1___closed__2; @@ -279,7 +277,6 @@ lean_object* l_Lean_HashMap_insert___at_Lean_Meta_AbstractMVars_abstractExprMVar lean_object* l_Lean_MetavarContext_getLevelDepth(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryResolve___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_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__2; -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__4; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__3___closed__2; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__8; @@ -292,7 +289,6 @@ LEAN_EXPORT uint8_t l_Lean_Meta_SynthInstance_Waiter_isRoot(lean_object*); lean_object* l_Lean_registerTraceClass(lean_object*, uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_consume___spec__4(lean_object*, lean_object*, size_t, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___spec__8(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_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___lambda__1___boxed(lean_object*, lean_object*); lean_object* l___private_Lean_Meta_Basic_0__Lean_Meta_withLocalDeclImp___rarg(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__1; @@ -311,7 +307,6 @@ lean_object* l_Lean_Meta_check(lean_object*, lean_object*, lean_object*, lean_ob static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_82____closed__1; static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__1___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__7; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__2; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__5; static lean_object* l_Lean_Meta_SynthInstance_generate___lambda__4___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getSubgoals___closed__1; @@ -333,6 +328,7 @@ static lean_object* l_Lean_Meta_SynthInstance_checkSystem___closed__2; lean_object* lean_array_pop(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__8; LEAN_EXPORT lean_object* lean_synth_pending(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); 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*); @@ -357,7 +353,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___rarg___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Util_Trace_0__Lean_getResetTraces___at_Lean_Meta_SynthInstance_newSubgoal___spec__11___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__3; LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Meta_SynthInstance_getInstances___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_levelZero; @@ -376,10 +371,10 @@ static lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_synthInstance LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_SynthInstance_newSubgoal___spec__2(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_synthPendingImp___lambda__2___closed__4; lean_object* l_Lean_Expr_constName_x3f(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__4; 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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__4; LEAN_EXPORT lean_object* l_Lean_addTrace___at_Lean_Meta_SynthInstance_wakeUp___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1; uint64_t l_Lean_Expr_hash(lean_object*); @@ -422,6 +417,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthI static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_6____closed__4; extern lean_object* l_Lean_trace_profiler_threshold; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_SynthInstance_State_tableEntries___default___spec__1(lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__11; static lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___closed__1; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, double, double, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -470,6 +466,7 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SynthInstance_get LEAN_EXPORT lean_object* l_MonadExcept_ofExcept___at_Lean_Meta_SynthInstance_generate___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_MVarId_isAssigned___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___rarg(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__9; lean_object* l_Lean_Meta_DiscrTree_getUnify___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer___closed__1; LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_tryResolve___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -485,7 +482,6 @@ LEAN_EXPORT lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSu LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKey___at_Lean_Meta_SynthInstance_mkTableKeyFor___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___lambda__2___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892_(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_instInhabitedAnswer___closed__2; extern lean_object* l_Lean_inheritedTraceOptions; lean_object* l_Lean_MessageData_ofExpr(lean_object*); @@ -508,7 +504,6 @@ static lean_object* l_Lean_Meta_SynthInstance_newSubgoal___lambda__2___closed__1 extern lean_object* l_Lean_crossEmoji; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode; static lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer___lambda__2___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__9; 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*); static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__2___closed__2; lean_object* l_Lean_HashMapImp_find_x3f___at___private_Lean_Meta_AbstractMVars_0__Lean_Meta_AbstractMVars_abstractLevelMVars___spec__1(lean_object*, lean_object*); @@ -517,6 +512,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_MkTableKey_instMonadMCtxM___l LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_mkAnswer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_foldrMUnsafe_fold___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_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_EXPORT lean_object* l_Lean_withTraceNode_x27___at_Lean_Meta_SynthInstance_newSubgoal___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*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__3; static lean_object* l_Lean_Meta_synthInstance_x3f___lambda__6___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_resume___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*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_Waiter_isRoot___boxed(lean_object*); @@ -529,7 +525,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getNextToResume___rarg___boxe 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_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_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__3; static lean_object* l_Lean_Meta_synthInstance_x3f_assignOutParams___closed__5; LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getTop___boxed(lean_object*); @@ -545,6 +540,7 @@ lean_object* l_Lean_MVarId_getDecl(lean_object*, lean_object*, lean_object*, lea 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_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698_(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_SynthInstance_findEntry_x3f___spec__2___boxed(lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_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*); @@ -573,7 +569,6 @@ static lean_object* l_Lean_Meta_SynthInstance_addAnswer___lambda__4___closed__11 lean_object* l_Lean_indentExpr(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__13; lean_object* l_Lean_Core_checkMaxHeartbeatsCore(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_removeUnusedArguments_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Meta_SynthInstance_getInstances___spec__7___closed__2; @@ -599,7 +594,6 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_State_result_x3f___default; static lean_object* l_panic___at_Lean_Meta_SynthInstance_getEntry___spec__1___closed__1; static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__1; static lean_object* l_Lean_withTraceNode___at_Lean_Meta_SynthInstance_newSubgoal___spec__9___lambda__4___closed__3; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__11; lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_Profile___hyg_6____spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_synthInstance_maxHeartbeats; @@ -654,6 +648,7 @@ lean_object* l_Lean_exceptOptionEmoji___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_backward_synthInstance_canonInstances; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_newSubgoal___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_withTraceNode___at_Lean_Meta_synthInstance_x3f___spec__13(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_wakeUp(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Level_param___override(lean_object*); size_t lean_usize_add(size_t, size_t); @@ -714,12 +709,10 @@ LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_SynthInstance_newSubgoal_ LEAN_EXPORT lean_object* l_Lean_Loop_forIn_loop___at_Lean_Meta_SynthInstance_getSubgoals___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_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_getSubgoals(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_SynthInstance_removeUnusedArguments_x3f___spec__2(lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); uint8_t l_Array_contains___at___private_Lean_Meta_FunInfo_0__Lean_Meta_collectDeps_visit___spec__2(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_isEqvAux___at_Lean_Meta_synthInstance_x3f___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__12; 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; @@ -757,6 +750,7 @@ lean_object* lean_expr_instantiate1(lean_object*, lean_object*); lean_object* l___private_Init_Data_Repr_0__Nat_reprFast(lean_object*); static lean_object* l_Lean_Meta_SynthInstance_getInstances___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_mkTableKeyFor___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__1; static lean_object* l_Lean_Meta_SynthInstance_tryResolve___lambda__1___closed__4; static lean_object* l_Lean_Meta_SynthInstance_instInhabitedGeneratorNode___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_SynthInstance_tryAnswer___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -22125,205 +22119,6 @@ x_8 = l_Lean_Meta_forallTelescopeReducing___at_Lean_Meta_getParamNames___spec__2 return x_8; } } -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -if (lean_obj_tag(x_1) == 0) -{ -lean_object* x_8; -x_8 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_8, 0, x_2); -lean_ctor_set(x_8, 1, x_7); -return x_8; -} -else -{ -lean_object* x_9; lean_object* x_10; uint8_t x_11; -x_9 = lean_ctor_get(x_1, 0); -lean_inc(x_9); -x_10 = lean_ctor_get(x_1, 1); -lean_inc(x_10); -lean_dec(x_1); -x_11 = !lean_is_exclusive(x_2); -if (x_11 == 0) -{ -lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; -x_12 = lean_ctor_get(x_2, 0); -x_13 = lean_ctor_get(x_2, 1); -x_14 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); -x_15 = lean_ctor_get(x_14, 0); -lean_inc(x_15); -x_16 = lean_ctor_get(x_14, 1); -lean_inc(x_16); -lean_dec(x_14); -x_17 = l_Lean_Level_hasMVar(x_15); -if (x_17 == 0) -{ -lean_object* x_18; -x_18 = lean_array_push(x_13, x_15); -lean_ctor_set(x_2, 1, x_18); -x_1 = x_10; -x_7 = x_16; -goto _start; -} -else -{ -lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; uint8_t x_24; lean_object* x_25; -lean_dec(x_15); -lean_dec(x_12); -x_20 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, x_16); -x_21 = lean_ctor_get(x_20, 0); -lean_inc(x_21); -x_22 = lean_ctor_get(x_20, 1); -lean_inc(x_22); -lean_dec(x_20); -x_23 = lean_array_push(x_13, x_21); -x_24 = 1; -x_25 = lean_box(x_24); -lean_ctor_set(x_2, 1, x_23); -lean_ctor_set(x_2, 0, x_25); -x_1 = x_10; -x_7 = x_22; -goto _start; -} -} -else -{ -lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; uint8_t x_32; -x_27 = lean_ctor_get(x_2, 0); -x_28 = lean_ctor_get(x_2, 1); -lean_inc(x_28); -lean_inc(x_27); -lean_dec(x_2); -x_29 = l_Lean_instantiateLevelMVars___at_Lean_Meta_normalizeLevel___spec__1(x_9, x_3, x_4, x_5, x_6, x_7); -x_30 = lean_ctor_get(x_29, 0); -lean_inc(x_30); -x_31 = lean_ctor_get(x_29, 1); -lean_inc(x_31); -lean_dec(x_29); -x_32 = l_Lean_Level_hasMVar(x_30); -if (x_32 == 0) -{ -lean_object* x_33; lean_object* x_34; -x_33 = lean_array_push(x_28, x_30); -x_34 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_34, 0, x_27); -lean_ctor_set(x_34, 1, x_33); -x_1 = x_10; -x_2 = x_34; -x_7 = x_31; -goto _start; -} -else -{ -lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; uint8_t x_40; lean_object* x_41; lean_object* x_42; -lean_dec(x_30); -lean_dec(x_27); -x_36 = l_Lean_Meta_mkFreshLevelMVar(x_3, x_4, x_5, x_6, x_31); -x_37 = lean_ctor_get(x_36, 0); -lean_inc(x_37); -x_38 = lean_ctor_get(x_36, 1); -lean_inc(x_38); -lean_dec(x_36); -x_39 = lean_array_push(x_28, x_37); -x_40 = 1; -x_41 = lean_box(x_40); -x_42 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_42, 0, x_41); -lean_ctor_set(x_42, 1, x_39); -x_1 = x_10; -x_2 = x_42; -x_7 = x_38; -goto _start; -} -} -} -} -} -static lean_object* _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1() { -_start: -{ -uint8_t x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = 0; -x_2 = l_Lean_Meta_SynthInstance_instInhabitedInstance___closed__2; -x_3 = lean_box(x_1); -x_4 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_4, 0, x_3); -lean_ctor_set(x_4, 1, x_2); -return x_4; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; lean_object* x_8; uint8_t x_9; -x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1; -x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_7, x_2, x_3, x_4, x_5, x_6); -x_9 = !lean_is_exclusive(x_8); -if (x_9 == 0) -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; -x_10 = lean_ctor_get(x_8, 0); -x_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); -lean_inc(x_12); -lean_dec(x_10); -x_13 = lean_array_to_list(lean_box(0), x_12); -x_14 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_14, 0, x_13); -lean_ctor_set(x_14, 1, x_11); -lean_ctor_set(x_8, 0, x_14); -return x_8; -} -else -{ -lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; -x_15 = lean_ctor_get(x_8, 0); -x_16 = lean_ctor_get(x_8, 1); -lean_inc(x_16); -lean_inc(x_15); -lean_dec(x_8); -x_17 = lean_ctor_get(x_15, 0); -lean_inc(x_17); -x_18 = lean_ctor_get(x_15, 1); -lean_inc(x_18); -lean_dec(x_15); -x_19 = lean_array_to_list(lean_box(0), x_18); -x_20 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_20, 0, x_19); -lean_ctor_set(x_20, 1, x_17); -x_21 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_21, 0, x_20); -lean_ctor_set(x_21, 1, x_16); -return x_21; -} -} -} -LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { -_start: -{ -lean_object* x_8; -x_8 = l_List_forIn_loop___at___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7); -lean_dec(x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -return x_8; -} -} -LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { -_start: -{ -lean_object* x_7; -x_7 = l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels(x_1, x_2, x_3, x_4, x_5, x_6); -lean_dec(x_5); -lean_dec(x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_7; -} -} LEAN_EXPORT lean_object* l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8, lean_object* x_9, lean_object* x_10) { _start: { @@ -29386,7 +29181,7 @@ lean_dec(x_3); return x_9; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -29396,17 +29191,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_9892____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__1; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__3() { _start: { lean_object* x_1; @@ -29414,17 +29209,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_9892____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__2; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__3; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__2; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__5() { _start: { lean_object* x_1; @@ -29432,37 +29227,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_9892____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__4; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__5; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__7() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__6; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__8() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__7; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__9() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__9() { _start: { lean_object* x_1; @@ -29470,17 +29265,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_9892____closed__10() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__8; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__9; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__8; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__11() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__11() { _start: { lean_object* x_1; @@ -29488,33 +29283,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_9892____closed__12() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__10; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__11; +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__10; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__13() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____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_9892____closed__12; -x_2 = lean_unsigned_to_nat(9892u); +x_1 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__12; +x_2 = lean_unsigned_to_nat(9698u); 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_9892_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698_(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_9892____closed__13; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__13; x_5 = l_Lean_registerTraceClass(x_2, x_3, x_4, x_1); if (lean_obj_tag(x_5) == 0) { @@ -30083,8 +29878,6 @@ l_Lean_Meta_SynthInstance_main___lambda__1___closed__5 = _init_l_Lean_Meta_Synth lean_mark_persistent(l_Lean_Meta_SynthInstance_main___lambda__1___closed__5); l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1 = _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1(); lean_mark_persistent(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocess___closed__1); -l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1 = _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1(); -lean_mark_persistent(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessLevels___closed__1); l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1 = _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1(); lean_mark_persistent(l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__1); l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2 = _init_l___private_Lean_Meta_SynthInstance_0__Lean_Meta_preprocessArgs___closed__2(); @@ -30157,33 +29950,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_9892____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__6); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__7(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__7); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__8(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__8); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__9(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__9); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__10(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__10); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__11(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__11); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__12(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__12); -l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__13(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892____closed__13); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9892_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__6); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__7 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__7(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__7); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__8 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__8(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__8); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__9 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__9(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__9); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__10 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__10(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__10); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__11 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__11(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__11); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__12 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__12(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__12); +l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__13 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__13(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698____closed__13); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_SynthInstance___hyg_9698_(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/AuxLemma.c b/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c index e1922fd6f785..162459b3a495 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c +++ b/stage0/stdlib/Lean/Meta/Tactic/AuxLemma.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Meta.Tactic.AuxLemma -// Imports: Lean.Meta.Basic +// Imports: Lean.AddDecl Lean.Meta.Basic #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -52,6 +52,7 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkAuxLemma___closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_mkAuxLemma___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_mkAuxLemma___spec__5___closed__1; +uint8_t l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_343____spec__1(lean_object*, lean_object*); lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); @@ -68,7 +69,6 @@ lean_object* lean_usize_to_nat(size_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_mkAuxLemma___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_AuxLemmas_lemmas___default___closed__3; static lean_object* l_Lean_Meta_instInhabitedAuxLemmas___closed__1; -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*); static size_t l_Lean_PersistentHashMap_findAux___at_Lean_Meta_mkAuxLemma___spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_mkAuxLemma___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); @@ -1452,7 +1452,7 @@ lean_inc(x_94); x_95 = lean_ctor_get(x_93, 1); lean_inc(x_95); lean_dec(x_93); -x_96 = l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_39____spec__1(x_1, x_95); +x_96 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_343____spec__1(x_1, x_95); lean_dec(x_95); if (x_96 == 0) { @@ -1908,7 +1908,7 @@ lean_inc(x_216); x_217 = lean_ctor_get(x_215, 1); lean_inc(x_217); lean_dec(x_215); -x_218 = l_List_beq___at___private_Lean_Data_OpenDecl_0__Lean_beqOpenDecl____x40_Lean_Data_OpenDecl___hyg_39____spec__1(x_1, x_217); +x_218 = l_List_beq___at___private_Lean_Declaration_0__Lean_beqConstantVal____x40_Lean_Declaration___hyg_343____spec__1(x_1, x_217); lean_dec(x_217); if (x_218 == 0) { @@ -2136,12 +2136,16 @@ lean_dec(x_4); return x_9; } } +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Meta_Basic(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Meta_Tactic_AuxLemma(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_AddDecl(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); diff --git a/stage0/stdlib/Lean/Meta/Tactic/LibrarySearch.c b/stage0/stdlib/Lean/Meta/Tactic/LibrarySearch.c index 2200527e4d61..fc3fca5c17e0 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/LibrarySearch.c +++ b/stage0/stdlib/Lean/Meta/Tactic/LibrarySearch.c @@ -66,7 +66,6 @@ lean_object* l_Lean_Meta_forallTelescope___at_Lean_Meta_mapForallTelescope_x27__ static lean_object* l_Lean_Meta_LibrarySearch_abortSpeculation___rarg___closed__2; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_LibrarySearch_interleaveWith___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_LibrarySearch_0__Lean_Meta_LibrarySearch_librarySearchLemma___lambda__1___closed__6; -static lean_object* l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3; uint8_t lean_float_decLt(double, double); LEAN_EXPORT lean_object* l_Lean_Meta_LibrarySearch_abortSpeculation___at___private_Lean_Meta_Tactic_LibrarySearch_0__Lean_Meta_LibrarySearch_librarySearch_x27___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_LibrarySearch_isAbortSpeculation___boxed(lean_object*); @@ -2168,26 +2167,18 @@ static lean_object* _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tact _start: { lean_object* x_1; -x_1 = lean_mk_string_from_bytes("Std", 3); -return x_1; -} -} -static lean_object* _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2() { -_start: -{ -lean_object* x_1; x_1 = lean_mk_string_from_bytes("abortSpeculation", 16); return x_1; } } -static lean_object* _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3() { +static lean_object* _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____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_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__1; -x_2 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_5____closed__1; +x_1 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_5____closed__4; +x_2 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_5____closed__6; x_3 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_5____closed__8; -x_4 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2; +x_4 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } @@ -2196,7 +2187,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tacti _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3; +x_2 = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2; x_3 = l_Lean_registerInternalExceptionId(x_2, x_1); return x_3; } @@ -6398,8 +6389,6 @@ l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_112 lean_mark_persistent(l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__1); l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2 = _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2(); lean_mark_persistent(l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__2); -l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3 = _init_l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3(); -lean_mark_persistent(l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120____closed__3); if (builtin) {res = l_Lean_Meta_LibrarySearch_initFn____x40_Lean_Meta_Tactic_LibrarySearch___hyg_1120_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l___private_Lean_Meta_Tactic_LibrarySearch_0__Lean_Meta_LibrarySearch_abortSpeculationId = lean_io_result_get_value(res); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c index c28d5831f58c..368bf4e02664 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Attr.c @@ -14,7 +14,6 @@ extern "C" { #endif static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__10; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__6; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__13; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__12; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__4; @@ -27,6 +26,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAtt lean_object* l_Lean_Attribute_erase(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__15; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__3; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__3; lean_object* lean_mk_empty_array_with_capacity(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__22; @@ -37,7 +37,6 @@ lean_object* l_Lean_Meta_isProp(lean_object*, lean_object*, lean_object*, lean_o LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_registerSimpAttr___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__9; lean_object* l_Array_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__3; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__15; lean_object* lean_array_push(lean_object*, lean_object*); lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); @@ -74,17 +73,21 @@ static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__5; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_registerSimpAttr___spec__5(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__1; size_t lean_usize_of_nat(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__3; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__14; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__12; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__16; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4; static lean_object* l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__3; lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_sevalSimpExtension; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__28; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__9; static lean_object* l_Lean_Meta_getSimpTheorems___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__1; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__15; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__7; lean_object* l_Lean_Meta_Origin_key(lean_object*); @@ -102,12 +105,12 @@ static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___sp static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__11; lean_object* lean_st_mk_ref(lean_object*, lean_object*); static lean_object* l_Lean_Meta_getSEvalTheorems___closed__1; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4; lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_isSimproc(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_registerSimpAttr___closed__1; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687_(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_mkSimpAttr___spec__2(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -117,13 +120,13 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Context_mkDefault(lean_object*, lean_o static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__25; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_simpExtension; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__3___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__11; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5; +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713_(lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__12; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__2; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -132,19 +135,19 @@ size_t lean_hashmap_mk_idx(lean_object*, uint64_t); static lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__3___closed__3; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_registerSimpAttr___spec__3(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__24; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2; lean_object* l_Lean_Meta_SimpExtension_getTheorems(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_addSimpTheorem(lean_object*, lean_object*, uint8_t, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__27; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__19; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpAttr___lambda__2___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_registerSimpAttr___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__4; lean_object* l_Lean_Meta_Simp_isBuiltinSimproc(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___spec__3___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_647_; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__1; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__14; @@ -162,14 +165,14 @@ extern lean_object* l_Lean_Expr_instHashable; static lean_object* l_Lean_Meta_mkSimpAttr___lambda__1___closed__9; lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_617_; lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__18; static lean_object* l_Lean_Meta_Simp_Context_mkDefault___rarg___closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__18; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__4; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Context_mkDefault___rarg___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__16; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -179,7 +182,6 @@ extern lean_object* l_Lean_Meta_Simp_defaultMaxSteps; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__23; extern lean_object* l_Lean_Meta_simpExtensionMapRef; lean_object* lean_array_uget(lean_object*, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747_(lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__6; lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -189,18 +191,16 @@ static lean_object* l_Lean_Meta_SimpTheorems_erase___at_Lean_Meta_mkSimpAttr___s static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__8; lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); lean_object* lean_array_get_size(lean_object*); -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_nat_add(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__5; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_getSimpTheorems(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773_(lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__17; -static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5; lean_object* l_Lean_Meta_SimpTheorems_eraseCore(lean_object*, lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__19; +static lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_mkSimpAttr___spec__1___closed__8; @@ -2249,7 +2249,7 @@ lean_dec(x_1); return x_7; } } -static lean_object* _init_l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_647_() { +static lean_object* _init_l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_617_() { _start: { lean_object* x_1; @@ -2698,7 +2698,7 @@ x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__1() { _start: { lean_object* x_1; @@ -2706,17 +2706,17 @@ x_1 = lean_mk_string_from_bytes("simp", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____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_Simp_Attr___hyg_747____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3() { _start: { lean_object* x_1; @@ -2724,7 +2724,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_Simp_Attr___hyg_747____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4() { _start: { lean_object* x_1; @@ -2732,18 +2732,18 @@ x_1 = lean_mk_string_from_bytes("simpExtension", 13); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4; 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_Simp_Attr___hyg_747____closed__6() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6() { _start: { lean_object* x_1; @@ -2751,18 +2751,18 @@ x_1 = lean_mk_string_from_bytes("simplification theorem", 22); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__6; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5; x_5 = l_Lean_Meta_registerSimpAttr(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__1() { _start: { lean_object* x_1; @@ -2770,17 +2770,17 @@ x_1 = lean_mk_string_from_bytes("seval", 5); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____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_Simp_Attr___hyg_773____closed__3() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__3() { _start: { lean_object* x_1; @@ -2788,18 +2788,18 @@ x_1 = lean_mk_string_from_bytes("sevalSimpExtension", 18); return x_1; } } -static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; x_1 = l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_8____closed__1; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__3; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____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_Simp_Attr___hyg_773____closed__5() { +static lean_object* _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5() { _start: { lean_object* x_1; @@ -2807,13 +2807,13 @@ x_1 = lean_mk_string_from_bytes("symbolic evaluator theorem", 26); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2; -x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__5; -x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4; +x_2 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2; +x_3 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5; +x_4 = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4; x_5 = l_Lean_Meta_registerSimpAttr(x_2, x_3, x_4, x_1); return x_5; } @@ -3170,38 +3170,38 @@ l_Lean_Meta_mkSimpAttr___lambda__1___closed__15 = _init_l_Lean_Meta_mkSimpAttr__ lean_mark_persistent(l_Lean_Meta_mkSimpAttr___lambda__1___closed__15); l_Lean_Meta_mkSimpAttr___lambda__1___closed__16 = _init_l_Lean_Meta_mkSimpAttr___lambda__1___closed__16(); lean_mark_persistent(l_Lean_Meta_mkSimpAttr___lambda__1___closed__16); -l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_647_ = _init_l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_647_(); -lean_mark_persistent(l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_647_); +l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_617_ = _init_l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_617_(); +lean_mark_persistent(l___auto____x40_Lean_Meta_Tactic_Simp_Attr___hyg_617_); l_Lean_Meta_registerSimpAttr___closed__1 = _init_l_Lean_Meta_registerSimpAttr___closed__1(); lean_mark_persistent(l_Lean_Meta_registerSimpAttr___closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__5); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__6(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747____closed__6); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_747_(lean_io_mk_world()); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__5); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687____closed__6); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_687_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_simpExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_simpExtension); lean_dec_ref(res); -}l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__1); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__2); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__3(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__3); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__4); -l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__5(); -lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773____closed__5); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_773_(lean_io_mk_world()); +}l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__1 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__1(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__1); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__2); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__3 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__3(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__3); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__4); +l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5 = _init_l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5(); +lean_mark_persistent(l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713____closed__5); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_Attr___hyg_713_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_sevalSimpExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_sevalSimpExtension); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c index 38fb57cb94a7..f4cbf78772e4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/SimpTheorems.c @@ -300,6 +300,7 @@ lean_object* l_Lean_MessageData_ofExpr(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4959____closed__12; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__6(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_SimpTheorems_erase___spec__2___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017_(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_checkBadRewrite___closed__1; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_4959____closed__25; @@ -310,12 +311,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_addSimpTheorem(lean_object*, lean_object*, static lean_object* l_Lean_Meta_isRflProofCore___closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_instLTOrigin; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__6; -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047____spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_addSimpTheoremEntry___spec__3(lean_object*, size_t, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Meta_SimpTheorems_isLemma___spec__2(lean_object*, size_t, lean_object*); static lean_object* l_Lean_Meta_ppSimpTheorem___rarg___lambda__1___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__3___closed__10; LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_shouldPreprocess(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017____spec__1(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__2___closed__1; LEAN_EXPORT lean_object* l_List_forIn_loop___at___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremsFromConst___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_mkSimpTheoremCore___lambda__2___closed__2; @@ -359,7 +360,6 @@ lean_object* l_id___rarg___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_addSimpTheoremEntry___spec__7___boxed(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_instInhabitedSimpTheorems; -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047_(lean_object*); lean_object* l_Lean_indentExpr(lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_SimpTheoremsArray_isLetDeclToUnfold___spec__1(lean_object*, lean_object*, size_t, size_t); @@ -415,7 +415,6 @@ lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_SimpTheoremsArray_eraseTheorem___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_List_reverse___rarg(lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Meta_addSimpTheorem___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_addSimpTheoremEntry___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_SimpTheorems_0__Lean_Meta_preprocess_go___lambda__8(lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -450,6 +449,7 @@ 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_Meta_Tactic_Simp_SimpTheorems___hyg_4959____closed__20; static lean_object* l_Lean_Meta_instToFormatSimpTheorem___closed__5; +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017____spec__1___boxed(lean_object*); lean_object* l_List_redLength___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at_Lean_Meta_SimpTheorems_isLemma___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_mkSimpTheorems(lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -13522,7 +13522,7 @@ x_7 = l_Lean_registerSimpleScopedEnvExtension___rarg(x_6, x_2); return x_7; } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017____spec__1(lean_object* x_1) { _start: { lean_object* x_2; @@ -13530,7 +13530,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -13557,11 +13557,11 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047____spec__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017____spec__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047____spec__1(x_1); +x_2 = l_Lean_mkHashMap___at_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017____spec__1(x_1); lean_dec(x_1); return x_2; } @@ -16396,7 +16396,7 @@ l_Lean_Meta_mkSimpExt___closed__2 = _init_l_Lean_Meta_mkSimpExt___closed__2(); lean_mark_persistent(l_Lean_Meta_mkSimpExt___closed__2); l_Lean_Meta_mkSimpExt___closed__3 = _init_l_Lean_Meta_mkSimpExt___closed__3(); lean_mark_persistent(l_Lean_Meta_mkSimpExt___closed__3); -if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5047_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Meta_initFn____x40_Lean_Meta_Tactic_Simp_SimpTheorems___hyg_5017_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_simpExtensionMapRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_simpExtensionMapRef); diff --git a/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c b/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c index 4e15e5867f3e..f29ab8ba32c4 100644 --- a/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c +++ b/stage0/stdlib/Lean/Meta/Tactic/Simp/Simproc.c @@ -21,18 +21,19 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocAttrCore___boxed(lean_object static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__15; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__5(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_Simprocs_add___spec__3___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_eraseSimprocAttr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_Simp_getSimprocExtensionCore_x3f___spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__14; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___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*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__6; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDeclExtState; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Meta_Simp_getSimprocFromDeclImpl___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_Meta_mkConstWithFreshMVarLevels___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__20; @@ -43,16 +44,18 @@ static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__5; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlM___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__5(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__4; size_t lean_usize_shift_right(size_t, size_t); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__9(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simprocCore___closed__6; static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__4; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__3; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__11; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__9; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___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*); lean_object* l_Lean_Meta_Simp_getDtConfig(lean_object*); @@ -64,7 +67,7 @@ uint8_t lean_usize_dec_le(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Simp_Simprocs_addCore___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__21; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__22; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocArrayCore(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocCore___boxed(lean_object*, 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_instBEq___rarg___boxed(lean_object*, lean_object*, lean_object*); @@ -72,7 +75,6 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Simp_ size_t lean_uint64_to_usize(uint64_t); static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_eraseSimprocAttr___lambda__2(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_Meta_Simp_dsimprocCore___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*); @@ -80,16 +82,18 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__3(lean_objec LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSEvalprocBuiltinAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2; lean_object* l_Array_qpartition___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Simp_Simprocs_addCore___spec__1___closed__3; lean_object* l_Lean_instBEqLocalInstance___boxed(lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_insert___at_Lean_NameSSet_insert___spec__2(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_DiscrTree_instInhabited(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore(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_mk_array(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocArrayCore___spec__1(uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___lambda__3___boxed(lean_object**); @@ -97,7 +101,6 @@ lean_object* l___private_Init_GetElem_0__outOfBounds___rarg(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocs___rarg___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simpAttrNameToSimprocAttrName___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__4(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_1155_(lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); @@ -110,7 +113,6 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_ static lean_object* l_Lean_Meta_Simp_registerSimproc___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__10(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Meta_Simp_getSimprocFromDeclImpl___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___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*); @@ -121,17 +123,14 @@ static lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDecl___closed__1; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__6___closed__1; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreDSimprocs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128_(lean_object*); static lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDecl___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_DStep_addExtraArgs(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simprocCore___closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore___boxed(lean_object*, lean_object*, lean_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_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5376____closed__4; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_add___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__11(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11; uint8_t l_Lean_Name_quickLt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__12___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_eraseSimprocAttr___closed__3; @@ -162,12 +161,10 @@ static lean_object* l_Lean_Meta_Simp_addSimprocAttrCore___closed__3; lean_object* l_instBEqProd___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocArrayCore___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_binSearchAux___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3; static lean_object* l_Lean_Meta_Simp_addSimprocAttrCore___closed__4; lean_object* l_Lean_Level_ofNat(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__10; lean_object* l_Lean_Expr_appArg_x21(lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__18; LEAN_EXPORT uint8_t l_Lean_Meta_Simp_SimprocDecl_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Init_Data_Array_BinSearch_0__Array_binInsertAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__12(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -175,18 +172,18 @@ static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__19; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__12; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__3; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5376____closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreDSimprocs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocAttr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocArrayCore(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_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_1121_(lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5; lean_object* l_instHashableArray___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_panic___at_Lean_Meta_Simp_Simprocs_addCore___spec__13(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostDSimprocs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocs; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___lambda__2(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*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__12; @@ -197,6 +194,8 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Si LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Simp_Simprocs_addCore___spec__11(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__2___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertVal_loop___at_Lean_Meta_Simp_Simprocs_addCore___spec__10(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocsArray_add___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocArrayCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__7; @@ -205,19 +204,20 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Meta_Simp_getSimprocFromDeclImp lean_object* l_Array_reverse___rarg(lean_object*); lean_object* l_Lean_throwError___at_Lean_Expr_abstractRangeM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttr___closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_builtinSimprocsRef; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__7(lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4; LEAN_EXPORT lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705_; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_BuiltinSimprocs_procs___default; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocArrayCore___spec__1(uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_SimprocEntry_try___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocDeclKeys_x3f(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDeclExtState___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isBuiltinSimproc___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1; static size_t l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostSimprocs___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_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -227,7 +227,7 @@ LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Simp_Simprocs_add___sp lean_object* l_Lean_registerPersistentEnvExtensionUnsafe___rarg(lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_getModuleEntries___rarg(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocExtension_getSimprocs(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_instBEqSimprocEntry(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__4; @@ -244,47 +244,47 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Si lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__11; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10; LEAN_EXPORT lean_object* l_Lean_AssocList_replace___at_Lean_Meta_Simp_registerSimprocAttr___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__10; lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_TransformStep_toStep(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__25; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocCore(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___private_Lean_CoreM_0__Lean_Core_mkFreshNameImp(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__2(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttr___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__1; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__24; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocsArray_isErased___boxed(lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__5; static lean_object* l_Lean_Meta_Simp_registerSimproc___closed__2; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__19; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostSimprocs___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__2; static lean_object* l_Lean_Meta_Simp_simprocCore___closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__3___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_Simp_registerSimprocAttr___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__9___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__5; lean_object* l_Lean_Meta_DiscrTree_empty(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5376_(lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostDSimprocs___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_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___lambda__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocAttr(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5; lean_object* l_Lean_throwError___at_Lean_registerTagAttribute___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_MetavarContext_getExprAssignmentDomain___spec__2___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__6___closed__1; LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20; lean_object* lean_st_mk_ref(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocs(lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__5; @@ -294,11 +294,13 @@ static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_ static lean_object* l_Lean_Meta_Simp_mkSimprocExt___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore___lambda__1(uint8_t, 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_Meta_Simp_instBEqSimprocEntry___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__2; lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Simp_Simprocs_addCore___spec__1___closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simprocCore___closed__1; lean_object* l_Lean_ScopedEnvExtension_modifyState___rarg(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_append(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___lambda__3(lean_object*, uint8_t, 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___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__26; @@ -308,25 +310,22 @@ lean_object* l_EStateM_pure___rarg(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttrCore___closed__2; uint8_t l_Lean_Meta_DiscrTree_Key_lt(lean_object*, lean_object*); lean_object* l_Lean_isTracingEnabledFor___at_Lean_Meta_Simp_congrArgs___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_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocDeclKeys_x3f___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__6; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__9(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_add(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2; static lean_object* l_Lean_Meta_Simp_getSimprocs___rarg___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___rarg(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986_(lean_object*); static lean_object* l_Lean_Meta_Simp_simpAttrNameToSimprocAttrName___closed__3; -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986____spec__1(lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__21; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__4; static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Simp_Simprocs_addCore___spec__1___closed__4; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1(uint8_t, lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098_(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__6; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__19; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__18; lean_object* l_Lean_instHashableLocalInstance___boxed(lean_object*); @@ -334,21 +333,19 @@ lean_object* l_instHashableProd___rarg___boxed(lean_object*, lean_object*, lean_ static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__2; static lean_object* l_Lean_Meta_Simp_SimprocDeclExtState_newEntries___default___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__3(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3; uint8_t lean_name_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostSimprocs(lean_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_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__3; LEAN_EXPORT lean_object* l_Std_Range_forIn_loop___at_Lean_Meta_Simp_SimprocEntry_try___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* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__4(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Meta_Simp_Simprocs_addCore___spec__2(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_BuiltinSimprocs_procs___default___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__12; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_BuiltinSimprocs_keys___default; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocEntry_tryD(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedSimprocDecl; @@ -361,19 +358,19 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Si static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__4; lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocCore___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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22; LEAN_EXPORT uint8_t l_Array_anyMUnsafe_any___at_Lean_Meta_Simp_SimprocsArray_isErased___spec__1(lean_object*, lean_object*, size_t, size_t); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__16; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__2(lean_object*, size_t, size_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_BuiltinSimprocs_keys___default___spec__1(lean_object*); static lean_object* l_Lean_Meta_Simp_SimprocDeclExtState_newEntries___default___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreSimprocs(lean_object*, 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_registerSimprocAttr(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Expr_appFn_x21(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616_(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__13; static lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___closed__1; @@ -392,8 +389,8 @@ lean_object* lean_usize_to_nat(size_t); size_t lean_hashmap_mk_idx(lean_object*, uint64_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocCore___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__2(lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_addCore(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__1(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__16; @@ -406,11 +403,8 @@ LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_Simp_registerBu static lean_object* l_Lean_Meta_DiscrTree_insertCore___at_Lean_Meta_Simp_Simprocs_addCore___spec__1___closed__1; static lean_object* l_Lean_Meta_Simp_registerSimprocAttr___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_erase(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Meta_mkEqTrans_x3f(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Meta_DiscrTree_0__Lean_Meta_DiscrTree_insertAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__9___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocCore___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -419,12 +413,10 @@ LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_getSimpro LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPostDSimprocs(lean_object*, 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_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__1(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__11; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_add___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3; static lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___lambda__2___closed__2; static lean_object* l_Lean_Meta_Simp_addSimprocAttrCore___closed__2; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__5; @@ -433,9 +425,11 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___lambda__1 lean_object* l_Lean_Meta_Simp_Step_addExtraArgs(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Meta_DiscrTree_instInhabitedKey; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__19; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__4___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__10; LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__6___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerBuiltinSimproc(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__9; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocsArray_erase(lean_object*, lean_object*); @@ -443,8 +437,9 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_isBuiltinSimproc(lean_object*, lean_ob uint8_t lean_nat_dec_eq(lean_object*, lean_object*); lean_object* l_Lean_mkApp3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_SimprocDeclExtState_newEntries___default___closed__3; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2; LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_Simp_registerSimprocAttr___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__10; uint8_t lean_nat_dec_lt(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl(lean_object*, lean_object*, lean_object*); @@ -465,6 +460,7 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_Meta_Simp_ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttrCore(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_Meta_Simp_SimprocsArray_isErased(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_builtinSimprocDeclsRef; static lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__3___closed__3; static lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttrCore___closed__1; @@ -481,41 +477,44 @@ lean_object* l_Lean_addTrace___at_Lean_Meta_Simp_congrArgs___spec__2(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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5376____closed__5; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541_(lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__3(lean_object*, lean_object*); lean_object* l_Array_insertAt_x21___rarg(lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__25; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10; lean_object* lean_nat_mul(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__16; lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____lambda__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__8; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__8; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2; lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocsArray_add(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simprocCore___closed__7; static lean_object* l_Lean_Meta_Simp_registerSimproc___lambda__3___closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocArrayCore___lambda__1(uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601_(lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__17; lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956____spec__1___boxed(lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__20; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_instInhabitedBuiltinSimprocs; LEAN_EXPORT lean_object* l_Lean_HashMapImp_moveEntries___at_Lean_Meta_Simp_registerSimprocAttr___spec__4(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocDeclKeys_x3f___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Environment_getModuleIdxFor_x3f(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSEvalSimprocs___boxed(lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7; lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19; lean_object* l_Array_back___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_SimprocDecl_lt___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_eraseSimprocAttr___lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_SimprocEntry_tryD___closed__1; static lean_object* l_Lean_Meta_Simp_simpAttrNameToSimprocAttrName___closed__2; +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956____spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_dsimprocCore___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_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__7; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_Simprocs_addCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -525,13 +524,13 @@ LEAN_EXPORT lean_object* l_Array_binInsertM___at_Lean_Meta_Simp_Simprocs_addCore static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__1; size_t lean_usize_sub(size_t, size_t); lean_object* l_Lean_mkHashMapImp___rarg(lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___lambda__3(lean_object*, uint8_t, lean_object*, lean_object*, uint8_t, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Simp_Simprocs_add___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___lambda__1___closed__8; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocBuiltinAttrCore___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_Meta_Simp_Simprocs_add___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__4(lean_object*, size_t, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); size_t lean_usize_add(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_addSimprocAttrCore(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_Simp_SimprocsArray_erase___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -540,12 +539,12 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocExtensionMapRef; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_mkSimprocExt___lambda__2___boxed(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___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_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__6; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__13; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_uget(lean_object*, size_t); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__24; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Meta_Simp_getSimprocFromDeclImpl___spec__1(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5997_; static lean_object* l_Lean_Meta_Simp_instInhabitedBuiltinSimprocs___closed__2; LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_Simp_Simprocs_add___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_dsimprocCore___closed__1; @@ -567,6 +566,7 @@ LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreSimprocs___lambda__1(lean_objec static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__7; static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_1121____closed__2; LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Meta_Simp_isBuiltinSimproc___spec__2___boxed(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2; lean_object* lean_string_append(lean_object*, lean_object*); uint8_t l_Lean_PersistentHashMap_contains___at_Lean_NameSSet_contains___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -574,29 +574,25 @@ static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____ static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__23; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__3; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocSEvalExtension; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15; static lean_object* l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__2; lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); static lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___lambda__2___closed__1; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__14; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__4___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__17; -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676_(lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_find_x3f___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1; static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__18; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_builtinSEvalprocsRef; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Meta_Simp_Simprocs_addCore___spec__5(lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986____spec__1___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_Meta_Simp_registerSimprocAttr___spec__2___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_simprocCore___closed__2; uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_Simp_isBuiltinSimproc___spec__2(lean_object*, lean_object*); lean_object* l_Lean_Meta_Simp_getConfig___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6027_; lean_object* l_Lean_PersistentHashMap_erase___at_Lean_Meta_addInstanceEntry___spec__18(lean_object*, lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_getSimprocExtension_x3f(lean_object*, lean_object*); @@ -607,13 +603,13 @@ static lean_object* l_Lean_Meta_Simp_eraseSimprocAttr___closed__1; static lean_object* l_Lean_Meta_Simp_simprocCore___lambda__2___closed__1; static lean_object* l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__17; static lean_object* l_Lean_Meta_Simp_SimprocsArray_add___closed__1; -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2; uint8_t l_Lean_Exception_isRuntime(lean_object*); -static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_Meta_Simp_Simprocs_addCore___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_registerBuiltinSimprocCore___closed__2; static lean_object* l_Lean_Meta_Simp_addSimprocAttr___closed__6; extern lean_object* l_Lean_PersistentHashMap_empty___at_Lean_Meta_Instances_erased___default___spec__1; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5; LEAN_EXPORT lean_object* l_Lean_HashMapImp_contains___at_Lean_Meta_Simp_isBuiltinSimproc___spec__1___boxed(lean_object*, lean_object*); lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_BuiltinSimprocs_keys___default___spec__1___boxed(lean_object*); @@ -628,26 +624,30 @@ static lean_object* l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__2; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); LEAN_EXPORT uint8_t l_Lean_AssocList_contains___at_Lean_Meta_Simp_registerBuiltinSimprocCore___spec__7(lean_object*, lean_object*); +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocDeclExt; LEAN_EXPORT lean_object* l_Lean_Meta_Simp_userPreSimprocs___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_array_get(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Meta_Simp_eraseSimprocAttr___closed__2; LEAN_EXPORT lean_object* l_Array_binSearchAux___at_Lean_Meta_Simp_getSimprocDeclKeys_x3f___spec__6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_qsort_sort___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____spec__6(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_foldlM___at_Lean_Meta_Simp_registerSimprocAttr___spec__5(lean_object*, lean_object*); uint64_t l_Lean_Meta_DiscrTree_Key_hash(lean_object*); static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__13; LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_Meta_Simp_registerSimprocAttr___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Simp_Simprocs_add___spec__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simprocExtension; +static lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_Simp_simpAttrNameToSimprocAttrName(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Meta_Simp_Simprocs_add___spec__1(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_Array_forInUnsafe_loop___at_Lean_Meta_Simp_dsimprocCore___spec__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___boxed(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_Meta_Simp_addSimprocAttr___closed__15; +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5705____closed__1; LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_BuiltinSimprocs_keys___default___spec__1(lean_object* x_1) { _start: @@ -16507,7 +16507,7 @@ x_11 = l_Lean_registerBuiltinAttribute(x_10, x_5); return x_11; } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986____spec__1(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956____spec__1(lean_object* x_1) { _start: { lean_object* x_2; @@ -16515,7 +16515,7 @@ x_2 = l_Lean_mkHashMapImp___rarg(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; uint8_t x_5; @@ -16542,16 +16542,16 @@ return x_8; } } } -LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986____spec__1___boxed(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956____spec__1___boxed(lean_object* x_1) { _start: { lean_object* x_2; -x_2 = l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986____spec__1(x_1); +x_2 = l_Lean_mkHashMap___at_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956____spec__1(x_1); lean_dec(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6027_() { +static lean_object* _init_l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5997_() { _start: { lean_object* x_1; @@ -17000,7 +17000,7 @@ x_4 = lean_box(x_3); return x_4; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1() { _start: { lean_object* x_1; @@ -17008,17 +17008,17 @@ x_1 = lean_mk_string_from_bytes("simprocAttr", 11); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -17028,7 +17028,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4() { _start: { lean_object* x_1; @@ -17036,19 +17036,19 @@ x_1 = lean_mk_string_from_bytes("simprocExtension", 16); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6() { _start: { lean_object* x_1; @@ -17056,19 +17056,19 @@ x_1 = lean_mk_string_from_bytes("Simplification procedure", 24); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068_(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_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3; -x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3; +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5; x_6 = l_Lean_Meta_Simp_registerSimprocAttr(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1() { _start: { lean_object* x_1; @@ -17076,17 +17076,17 @@ x_1 = lean_mk_string_from_bytes("sevalprocAttr", 13); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3() { _start: { lean_object* x_1; lean_object* x_2; @@ -17096,7 +17096,7 @@ lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4() { _start: { lean_object* x_1; @@ -17104,19 +17104,19 @@ x_1 = lean_mk_string_from_bytes("simprocSEvalExtension", 21); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6() { _start: { lean_object* x_1; @@ -17124,14 +17124,14 @@ x_1 = lean_mk_string_from_bytes("Symbolic evaluator procedure", 28); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098_(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_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3; -x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3; +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5; x_6 = l_Lean_Meta_Simp_registerSimprocAttr(x_2, x_3, x_4, x_5, x_1); return x_6; } @@ -18091,7 +18091,7 @@ lean_dec(x_2); return x_7; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -18099,28 +18099,28 @@ x_1 = lean_mk_string_from_bytes("addSimprocBuiltinAttr", 21); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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; -x_7 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2; +x_7 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2; x_8 = l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -18128,25 +18128,25 @@ x_1 = lean_mk_string_from_bytes("Not implemented yet, [-builtin_simproc]", 39); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2; +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; @@ -18156,27 +18156,27 @@ x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4() { _start: { lean_object* x_1; @@ -18184,17 +18184,17 @@ x_1 = lean_mk_string_from_bytes("initFn", 6); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6() { _start: { lean_object* x_1; @@ -18202,67 +18202,67 @@ x_1 = lean_mk_string_from_bytes("_@", 2); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9; x_2 = l_Array_forInUnsafe_loop___at_Lean_Meta_Simp_simprocCore___spec__1___closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11; x_2 = l_Lean_Meta_Simp_getSimprocFromDeclImpl___closed__5; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13() { _start: { lean_object* x_1; @@ -18270,27 +18270,27 @@ x_1 = lean_mk_string_from_bytes("_hyg", 4); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14; -x_2 = lean_unsigned_to_nat(6601u); +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14; +x_2 = lean_unsigned_to_nat(6541u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16() { _start: { lean_object* x_1; @@ -18298,17 +18298,17 @@ x_1 = lean_mk_string_from_bytes("simprocBuiltinAttr", 18); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18() { _start: { lean_object* x_1; @@ -18316,13 +18316,13 @@ x_1 = lean_mk_string_from_bytes("Builtin simplification procedure", 32); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__17; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -18332,29 +18332,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__21() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__19; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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); @@ -18362,40 +18362,40 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____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; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -18403,28 +18403,28 @@ x_1 = lean_mk_string_from_bytes("addSEvalprocBuiltinAttr", 23); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__1; x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__2; x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_179____closed__3; -x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1; +x_4 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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; -x_7 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2; +x_7 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2; x_8 = l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin(x_1, x_2, x_7, x_4, x_5, x_6); return x_8; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1() { _start: { lean_object* x_1; @@ -18432,35 +18432,35 @@ x_1 = lean_mk_string_from_bytes("Not implemented yet, [-builtin_sevalproc]", 41) return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1; x_2 = l_Lean_stringToMessageData(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2; +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14; -x_2 = lean_unsigned_to_nat(6676u); +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14; +x_2 = lean_unsigned_to_nat(6616u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2() { _start: { lean_object* x_1; @@ -18468,17 +18468,17 @@ x_1 = lean_mk_string_from_bytes("sevalprocBuiltinAttr", 20); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2; x_3 = l_Lean_Name_str___override(x_1, x_2); return x_3; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4() { _start: { lean_object* x_1; @@ -18486,13 +18486,13 @@ x_1 = lean_mk_string_from_bytes("Builtin symbolic evaluation procedure", 37); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__3; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4; x_4 = 1; x_5 = lean_alloc_ctor(0, 3, 1); lean_ctor_set(x_5, 0, x_1); @@ -18502,29 +18502,29 @@ lean_ctor_set_uint8(x_5, sizeof(void*)*3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___boxed), 6, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___boxed), 6, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___boxed), 4, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___boxed), 4, 0); return x_1; } } -static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8() { +static lean_object* _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6; -x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7; +x_1 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__5; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6; +x_3 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7; x_4 = lean_alloc_ctor(0, 3, 0); lean_ctor_set(x_4, 0, x_1); lean_ctor_set(x_4, 1, x_2); @@ -18532,33 +18532,33 @@ lean_ctor_set(x_4, 2, x_3); return x_4; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8; +x_2 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8; x_3 = l_Lean_registerBuiltinAttribute(x_2, x_1); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); +x_8 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1(x_1, x_2, x_7, x_4, x_5, x_6); lean_dec(x_5); lean_dec(x_4); lean_dec(x_2); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____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_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2(x_1, x_2, x_3, x_4); +x_5 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2(x_1, x_2, x_3, x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); @@ -18866,7 +18866,7 @@ else { lean_object* x_8; lean_dec(x_1); -x_8 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2; +x_8 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2; return x_8; } } @@ -18874,7 +18874,7 @@ else { lean_object* x_9; lean_dec(x_1); -x_9 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2; +x_9 = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2; return x_9; } } @@ -19358,45 +19358,45 @@ l_Lean_Meta_Simp_addSimprocAttr___closed__24 = _init_l_Lean_Meta_Simp_addSimproc lean_mark_persistent(l_Lean_Meta_Simp_addSimprocAttr___closed__24); l_Lean_Meta_Simp_addSimprocAttr___closed__25 = _init_l_Lean_Meta_Simp_addSimprocAttr___closed__25(); lean_mark_persistent(l_Lean_Meta_Simp_addSimprocAttr___closed__25); -if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5986_(lean_io_mk_world()); +if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5956_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_Simp_simprocExtensionMapRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_Simp_simprocExtensionMapRef); lean_dec_ref(res); -}l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6027_ = _init_l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6027_(); -lean_mark_persistent(l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6027_); +}l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5997_ = _init_l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5997_(); +lean_mark_persistent(l___auto____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_5997_); l_Lean_Meta_Simp_registerSimprocAttr___closed__1 = _init_l_Lean_Meta_Simp_registerSimprocAttr___closed__1(); lean_mark_persistent(l_Lean_Meta_Simp_registerSimprocAttr___closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__3); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__4); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__5); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128____closed__6); -if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6128_(lean_io_mk_world()); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__3); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__4); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__5); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068____closed__6); +if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6068_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_Simp_simprocExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_Simp_simprocExtension); lean_dec_ref(res); -}l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__3); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__4); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__5); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158____closed__6); -if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6158_(lean_io_mk_world()); +}l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__3); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__4); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__5); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098____closed__6); +if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6098_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_Meta_Simp_simprocSEvalExtension = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_Meta_Simp_simprocSEvalExtension); @@ -19461,86 +19461,86 @@ l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed_ lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__19); l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__20 = _init_l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__20(); lean_mark_persistent(l___private_Lean_Meta_Tactic_Simp_Simproc_0__Lean_Meta_Simp_addBuiltin___closed__20); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__1___closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____lambda__2___closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__3); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__4); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__5); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__6); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__7); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__8); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__9); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__10); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__11); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__12); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__13); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__14); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__15); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__16); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__17); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__18); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__19); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__20); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__21); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601____closed__22); -if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6601_(lean_io_mk_world()); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__1___closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____lambda__2___closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__3); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__4); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__5); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__6); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__7); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__8); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__9); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__10); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__11); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__12); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__13); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__14); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__15); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__16); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__17 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__17(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__17); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__18); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__19 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__19(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__19); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__20); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__21 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__21(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__21); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541____closed__22); +if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6541_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__1___closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____lambda__2___closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__1); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__2); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__3); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__4); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__5); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__6); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__7); -l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8(); -lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676____closed__8); -if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6676_(lean_io_mk_world()); +}l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__1___closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____lambda__2___closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__1); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__2); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__3 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__3(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__3); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__4); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__5 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__5(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__5); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__6); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__7); +l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8 = _init_l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8(); +lean_mark_persistent(l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616____closed__8); +if (builtin) {res = l_Lean_Meta_Simp_initFn____x40_Lean_Meta_Tactic_Simp_Simproc___hyg_6616_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Meta_Simp_getSimprocs___rarg___closed__1 = _init_l_Lean_Meta_Simp_getSimprocs___rarg___closed__1(); diff --git a/stage0/stdlib/Lean/Meta/Transform.c b/stage0/stdlib/Lean/Meta/Transform.c index 87c62183ab06..0816b1489004 100644 --- a/stage0/stdlib/Lean/Meta/Transform.c +++ b/stage0/stdlib/Lean/Meta/Transform.c @@ -151,7 +151,6 @@ LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_transform_visit__ LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_transform_visit_visitForall___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Meta_transform_visit___spec__2___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*); lean_object* l_Lean_Core_instantiateValueLevelParams(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_withLocalDecl___at_Lean_Meta_transform_visit_visitLambda___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit___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* l_StateRefT_x27_lift___rarg___boxed(lean_object*, lean_object*); @@ -324,6 +323,7 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_ReaderT_bind___at_Lean_Meta_zetaReduce___spec__14___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitForall___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_ReaderT_bind___at_Lean_Core_transform_visit___spec__3___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Core_withIncRecDepth___at_Lean_Core_betaReduce___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Meta_transform_visit_visitLambda___at_Lean_Meta_zetaReduce___spec__4(lean_object*, lean_object*, uint8_t, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -10125,7 +10125,7 @@ x_13 = lean_ctor_get(x_11, 0); lean_inc(x_13); lean_dec(x_11); lean_inc(x_1); -x_14 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_1, x_3, x_4, x_12); +x_14 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_1, x_3, x_4, x_12); x_15 = lean_ctor_get(x_14, 1); lean_inc(x_15); lean_dec(x_14); @@ -10144,7 +10144,7 @@ lean_inc(x_19); x_20 = lean_ctor_get(x_18, 1); lean_inc(x_20); lean_dec(x_18); -x_21 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_9, x_3, x_4, x_20); +x_21 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_9, x_3, x_4, x_20); lean_dec(x_4); lean_dec(x_3); x_22 = !lean_is_exclusive(x_21); @@ -10176,7 +10176,7 @@ lean_inc(x_26); x_27 = lean_ctor_get(x_18, 1); lean_inc(x_27); lean_dec(x_18); -x_28 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_9, x_3, x_4, x_27); +x_28 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_9, x_3, x_4, x_27); lean_dec(x_4); lean_dec(x_3); x_29 = !lean_is_exclusive(x_28); diff --git a/stage0/stdlib/Lean/Parser/Extension.c b/stage0/stdlib/Lean/Parser/Extension.c index 3346b9906a50..c7e167883857 100644 --- a/stage0/stdlib/Lean/Parser/Extension.c +++ b/stage0/stdlib/Lean/Parser/Extension.c @@ -13,7 +13,6 @@ #ifdef __cplusplus extern "C" { #endif -LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489_(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__5; lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__14; @@ -21,12 +20,10 @@ LEAN_EXPORT lean_object* l_Lean_Parser_addTrailingParser(lean_object*, lean_obje LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_evalInsideQuot(lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4; lean_object* l_Lean_Name_reprPrec(lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5829_; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__17; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024_(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__2___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__10; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -43,12 +40,12 @@ size_t lean_usize_shift_right(size_t, size_t); lean_object* l_Lean_ScopedEnvExtension_activateScoped___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_Entry_toOLeanEntry(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865_(lean_object*); static lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3773____closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore(lean_object*); extern lean_object* l_Lean_Parser_SyntaxStack_empty; -LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_40____closed__3; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__1___boxed(lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2___closed__2; @@ -58,13 +55,13 @@ static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__6; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_getParserAliasInfo___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_contains___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__1___boxed(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4229____closed__3; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__6(size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ConstantInfo_type(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4516____closed__2; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3; LEAN_EXPORT uint8_t l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_194_(lean_object*); uint8_t lean_usize_dec_le(size_t, size_t); @@ -72,11 +69,11 @@ static lean_object* l_Lean_Parser_declareLeadingBuiltinParser___closed__1; static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__8; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__4; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__2(uint8_t, lean_object*, size_t, size_t, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Parser_getCategory___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getCategory(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_formatStxAux(lean_object*, uint8_t, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___boxed(lean_object*); static lean_object* l_Lean_Parser_ParserExtension_instInhabitedEntry___closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__2; @@ -84,7 +81,6 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe size_t lean_uint64_to_usize(uint64_t); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAux___at_Lean_Parser_getCategory___spec__2(lean_object*, size_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addParserTokens(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__6(lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__2___closed__2; lean_object* l_Lean_Data_Trie_empty(lean_object*); @@ -92,18 +88,17 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698 LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__3(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_Parser_builtinTokenTable; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1; lean_object* l_Lean_Parser_adaptCacheableContextFn(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_leadingNode(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Syntax_getId(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinTrailingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_categoryParser(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12; static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__17; lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1(lean_object*); static lean_object* l_Lean_Parser_getUnaryAlias___rarg___closed__1; -LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5645_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__2___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); size_t lean_usize_mul(size_t, size_t); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__1; @@ -111,7 +106,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_addToken(lean_object*, uint8_t, lean_obje LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__1___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_SyntaxStack_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserOfConstantUnsafe(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__4; +LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369_(lean_object*, lean_object*); uint8_t lean_usize_dec_eq(size_t, size_t); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4516_(lean_object*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Parser_withOpenDeclFnCore___spec__1___boxed(lean_object*, lean_object*, lean_object*); @@ -123,12 +118,13 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwP static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__11; lean_object* l_Lean_declareBuiltinDocStringAndRanges(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerParserAttributeHook(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4; lean_object* l_IO_ofExcept___at_Lean_declareBuiltin___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_nonReservedSymbol(lean_object*, uint8_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1; LEAN_EXPORT lean_object* l_List_forM___at_Lean_Parser_runParserAttributeHooks___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_array_fget(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStack___lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3; lean_object* lean_array_fset(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getParserAliasInfo(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__2; @@ -141,21 +137,22 @@ LEAN_EXPORT lean_object* l_Lean_Parser_withOpenFn(lean_object*, lean_object*, le static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__5; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__19; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__13; lean_object* lean_environment_find(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2; lean_object* l_Lean_Attribute_Builtin_ensureNoArgs(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__3(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerParserCategory(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3; static lean_object* l_Lean_Parser_getParserPriority___closed__4; static lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__15; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1; lean_object* l_Lean_PersistentHashMap_foldlMAux___at_Lean_mkModuleData___spec__3___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Parser_evalParserConstUnsafe___spec__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___closed__1; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserExtension_instInhabitedState___closed__1; @@ -166,18 +163,15 @@ LEAN_EXPORT lean_object* l_Lean_Parser_isValidSyntaxNodeKind___boxed(lean_object static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxKindOfParserAlias_x3f___boxed(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__21; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3773____lambda__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); uint8_t lean_string_dec_eq(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__22; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946_(lean_object*); lean_object* l_Lean_ScopedEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerScopedEnvExtensionUnsafe___rarg(lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_stackSize(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__7(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5; static lean_object* l_Lean_Parser_isParserCategory___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3773____closed__3; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__2(lean_object*, size_t, lean_object*); @@ -198,6 +192,7 @@ lean_object* l_Lean_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_swap(lean_object*, lean_object*, lean_object*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2(lean_object*, size_t, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3; static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__6; static lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2___closed__1; @@ -205,13 +200,13 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4229 lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__3; lean_object* l_Lean_PersistentHashMap_insert___at_Lean_Parser_SyntaxNodeKindSet_insert___spec__1(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); size_t lean_usize_of_nat(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserContext_resolveParserName(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_builtinParserCategoriesRef; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserState___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addParser(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__8; LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__5(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerBuiltinNodeKind(lean_object*, lean_object*); @@ -227,13 +222,12 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default; static lean_object* l_Lean_Parser_getBinaryAlias___rarg___closed__1; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__2; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1; lean_object* l_Lean_declareBuiltin(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7; static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2___boxed(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ResolveName_resolveNamespace(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getParserPriority___boxed(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985_(lean_object*); lean_object* l_List_eraseDups___at_Lean_ResolveName_resolveGlobalName_loop___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_Parser_getCategory___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_to_int(lean_object*); @@ -241,6 +235,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698 LEAN_EXPORT lean_object* l_Lean_Parser_getAlias___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_foldlM___at_Lean_Parser_addParserTokens___spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_withOpenFn___closed__1; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1(lean_object*, lean_object*); static lean_object* l_Lean_Parser_getParserAliasInfo___closed__1; static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__7; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__5; @@ -253,6 +248,7 @@ static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__ LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__5___rarg(lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__4___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_runParserCategory___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1; lean_object* lean_eval_const(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ensureBinaryParserAlias(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserAttributeImpl___closed__2; @@ -263,7 +259,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2(lean_o LEAN_EXPORT lean_object* l_Lean_Parser_builtinSyntaxNodeKindSetRef; lean_object* l_Lean_Parser_ParserState_toErrorMsg(lean_object*, lean_object*); static lean_object* l_Lean_Parser_ParserAliasInfo_stackSz_x3f___default___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__6; lean_object* l_Lean_Syntax_getKind(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameMap_contains___rarg(lean_object*, lean_object*); @@ -276,6 +272,7 @@ lean_object* l_Lean_registerAttributeOfBuilder(lean_object*, lean_object*, lean_ LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Parser_addToken___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_registerBuiltinParserAttribute___closed__2; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__8; lean_object* l_Lean_Parser_trailingNode(lean_object*, lean_object*, lean_object*, lean_object*); @@ -296,7 +293,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_parserExtension; LEAN_EXPORT lean_object* l_Lean_Parser_mkInputContext(lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12; lean_object* l_Lean_FileMap_ofString(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_Parser_sepBy1(lean_object*, lean_object*, lean_object*, uint8_t); @@ -305,27 +301,25 @@ LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_Parser lean_object* l_Lean_Parser_symbol(lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_leadingIdentBehavior___boxed(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2; lean_object* l_Lean_Name_num___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addParserCategory___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_filterMap___at_Lean_Parser_resolveParserNameCore___spec__2(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2; +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5; uint8_t lean_string_utf8_at_end(lean_object*, lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuiltinTokens___closed__1; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__7; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_declareTrailingBuiltinParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ensureConstantParserAlias(lean_object*, lean_object*); static lean_object* l_Lean_Parser_mkParserOfConstantUnsafe___closed__3; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_getSyntaxKindOfParserAlias_x3f___spec__1(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__4; +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5; lean_object* l_Lean_RBNode_insert___at_Lean_NameMap_insert___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__2; static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__4; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10; static lean_object* l_Lean_Parser_mkParserState___closed__1; -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752_(lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__11; LEAN_EXPORT lean_object* l_Lean_Parser_instCoeForallParserForallParserAliasValue(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTokenConfig___closed__2; @@ -334,12 +328,11 @@ uint8_t lean_name_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__3; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1; static lean_object* l_Lean_Parser_runParserCategory___closed__2; uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__1(lean_object*, lean_object*); lean_object* l_Lean_Parser_nodeWithAntiquot(lean_object*, lean_object*, lean_object*, uint8_t); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__5; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__9; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__1___rarg(lean_object*, lean_object*, lean_object*); @@ -347,8 +340,6 @@ lean_object* l_Lean_Syntax_getArg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux___spec__1(lean_object*, lean_object*); lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__1; uint8_t l___private_Init_Data_Option_Basic_0__Option_beqOption____x40_Init_Data_Option_Basic___hyg_158____at_Lean_Parser_ParserState_hasError___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_declName___default; @@ -365,14 +356,16 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_foldlMAux_traverse___at___priv LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr_visit(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAux___at_Lean_Parser_isValidSyntaxNodeKind___spec__2___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Parser_resolveParserNameCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9; LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2; static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__6; LEAN_EXPORT lean_object* l_Lean_Parser_declareLeadingBuiltinParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_ParserState_mkError(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_ParserExtension_OLeanEntry_toEntry(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_getAlias___spec__1___rarg___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__6; extern lean_object* l_Std_Format_defWidth; LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_instInhabitedOLeanEntry; @@ -385,6 +378,7 @@ LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184_; lean_object* l_Lean_KVMap_setBool(lean_object*, lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_parserAliases2infoRef; static lean_object* l_Lean_Parser_registerBuiltinParserAttribute___closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__8; lean_object* l_Lean_Data_Trie_find_x3f_loop___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__9; @@ -429,13 +423,14 @@ LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at___private_Lean_Par LEAN_EXPORT lean_object* l_Lean_Parser_getUnaryAlias___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2; LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_contains___at_Lean_Parser_isValidSyntaxNodeKind___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkParserAttributeImpl___elambda__2___boxed(lean_object*); static lean_object* l_Lean_Parser_registerBuiltinParserAttribute___closed__1; LEAN_EXPORT lean_object* l_Lean_Parser_compileParserDescr(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDeclFnCore(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_addEntryImpl(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____lambda__2___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_AttributeImpl_erase___default___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -452,43 +447,46 @@ LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__1; uint8_t lean_internal_is_stage0(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3603_(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4; static lean_object* l_Lean_Parser_withOpenDeclFnCore___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_40_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_2234_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getTokenTable___boxed(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__2___closed__4; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__8; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5739_; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_mkCategoryAntiquotParserFn(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_getSyntaxNodeKinds___closed__1; +LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5641_; static lean_object* l_Lean_Parser_parserOfStackFn___lambda__1___closed__3; LEAN_EXPORT lean_object* l_Lean_ofExcept___at_Lean_Parser_addToken___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3773_(lean_object*); uint8_t lean_nat_dec_lt(lean_object*, lean_object*); lean_object* l_Lean_Parser_withCache(lean_object*, lean_object*); lean_object* l_Lean_mkRawNatLit(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__3; +LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_registerAliasCore___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2; lean_object* l_id___rarg___boxed(lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____lambda__2___closed__1; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static size_t l_Lean_PersistentHashMap_containsAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__2___closed__1; static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__3; static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__3; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_mkParserOfConstant(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_getAlias___spec__1___rarg(lean_object*, lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStack(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntries(lean_object*, lean_object*); static lean_object* l_Lean_Parser_declareTrailingBuiltinParser___closed__1; static lean_object* l_Lean_Parser_resolveParserNameCore_isParser___closed__1; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addTokenConfig___closed__1; -static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5; -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Parser_registerAlias(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instCoeParserParserAliasValue(lean_object*); @@ -504,18 +502,21 @@ static lean_object* l_Lean_Parser_ParserExtension_addEntryImpl___closed__5; uint64_t l_Lean_Name_hash___override(lean_object*); static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined___rarg___closed__2; static uint8_t l_Lean_Parser_isValidSyntaxNodeKind___closed__1; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1; extern lean_object* l_Id_instMonad; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Parser_withOpenDeclFnCore___spec__1(size_t, size_t, lean_object*); lean_object* l_Repr_addAppParen(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_add___at_Lean_Parser_addToken___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_categoryParserFnImpl(lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_runParserAttributeHooks(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_getParserPriority___closed__5; LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias(lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at___private_Lean_Parser_Extension_0__Lean_Parser_addBuiltinParserCategory___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxNodeKinds(lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2; static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__5; lean_object* l_Lean_ResolveName_resolveGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); @@ -525,16 +526,15 @@ lean_object* l_Lean_Parser_prattParser(lean_object*, lean_object*, uint8_t, lean LEAN_EXPORT lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_containsAtAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_ScopedEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907_(lean_object*); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Parser_addLeadingParser___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStack___elambda__1(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_initCacheForInput(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_instReprParserName; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_getSyntaxKindOfParserAlias_x3f(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__5; static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___closed__4; lean_object* l_Lean_ScopedEnvExtension_addCore___rarg(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5615_; LEAN_EXPORT lean_object* l_Lean_Parser_withOpenDeclFn(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkCollisionNode___rarg(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_sepBy(lean_object*, lean_object*, lean_object*, uint8_t); @@ -544,13 +544,11 @@ static lean_object* l_Lean_Parser_getConstAlias___rarg___closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_isParserAlias___boxed(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_mkEmptyEntriesArray(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698_(lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_parserOfStackFn(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux___at___private_Lean_Parser_Extension_0__Lean_Parser_addParserCategoryCore___spec__5___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_erase_macro_scopes(lean_object*); lean_object* l_Lean_registerBuiltinAttribute(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1; lean_object* l_Lean_Parser_adaptUncacheableContextFn(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_ParserExtension_Entry_toOLeanEntry___boxed(lean_object*); static lean_object* l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__3; @@ -574,7 +572,6 @@ size_t lean_usize_add(size_t, size_t); static lean_object* l_Lean_Parser_instReprParserName___closed__1; lean_object* l_Lean_Parser_TokenMap_insert___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Parser_mkAntiquot(lean_object*, lean_object*, uint8_t, uint8_t); -LEAN_EXPORT lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5701_; lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_addTrailingParserAux___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___lambda__1(lean_object*, uint8_t, lean_object*); @@ -608,7 +605,6 @@ LEAN_EXPORT lean_object* l_Lean_Parser_ParserAliasInfo_autoGroupArgs___default__ LEAN_EXPORT lean_object* l_Lean_Parser_addParserCategory(lean_object*, lean_object*, lean_object*, uint8_t); static lean_object* l_Lean_Parser_withOpenFn___closed__2; LEAN_EXPORT uint8_t l_Lean_Parser_isValidSyntaxNodeKind(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2; LEAN_EXPORT lean_object* l_Lean_Parser_getBinaryAlias___rarg___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_mkCategoryAntiquotParser(lean_object*); @@ -616,7 +612,6 @@ lean_object* l_Lean_Syntax_isNatLit_x3f(lean_object*); static lean_object* l_Lean_Parser_parserOfStackFn___lambda__1___closed__1; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__10; lean_object* l_Lean_Parser_whitespace(lean_object*, lean_object*); -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4; LEAN_EXPORT lean_object* l_Lean_Parser_addBuiltinParser(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6____closed__1; uint8_t lean_nat_dec_le(lean_object*, lean_object*); @@ -626,6 +621,7 @@ uint8_t lean_usize_dec_lt(size_t, size_t); LEAN_EXPORT lean_object* l_List_foldl___at_Lean_Parser_addLeadingParser___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuiltinTokens(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_registerAttributeImplBuilder(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_4229_(lean_object*); lean_object* lean_nat_add(lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_getCollisionNodeSize___rarg(lean_object*); @@ -635,6 +631,7 @@ static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84__ uint8_t l_Lean_Exception_isRuntime(lean_object*); LEAN_EXPORT uint8_t l_Lean_Parser_leadingIdentBehavior(lean_object*, lean_object*); static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_84____closed__4; +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8; LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3773____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); extern lean_object* l_Lean_Parser_categoryParserFnRef; LEAN_EXPORT lean_object* l_Lean_Parser_resolveParserNameCore_isParser(lean_object*, lean_object*); @@ -646,6 +643,8 @@ LEAN_EXPORT lean_object* l_Lean_Parser_isParserAlias(lean_object*, lean_object*) LEAN_EXPORT lean_object* l_Lean_Parser_getConstAlias___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_addLeadingParser(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__24; +static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5; +static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__1; lean_object* lean_array_uset(lean_object*, size_t, lean_object*); lean_object* l_Lean_MessageData_ofName(lean_object*); @@ -663,16 +662,17 @@ static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_updateBuilt static lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_throwParserCategoryAlreadyDefined___rarg___closed__1; static lean_object* l_Lean_Parser_parserOfStackFn___closed__2; static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3993____closed__4; -static lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3; size_t lean_usize_land(size_t, size_t); LEAN_EXPORT lean_object* l_List_foldl___at___private_Lean_Parser_Extension_0__Lean_Parser_withNamespaces___spec__1(uint8_t, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_declareBuiltinParser___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Parser_categoryParserFnImpl___closed__5; +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662_(lean_object*); static lean_object* l___auto____x40_Lean_Parser_Extension___hyg_5184____closed__20; LEAN_EXPORT lean_object* l_Lean_RBNode_find___at_Lean_Parser_getSyntaxKindOfParserAlias_x3f___spec__1___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_BuiltinParserAttribute_add___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Parser_mkParserOfConstantUnsafe___spec__3___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_evalParserConstUnsafe(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904_(lean_object*); uint8_t l_Array_isEmpty___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_throwUnknownParserCategory___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Parser_withOpen(lean_object*); @@ -10243,7 +10243,7 @@ x_11 = lean_ctor_get(x_9, 1); lean_inc(x_11); lean_dec(x_9); x_12 = 0; -x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_5, x_12); +x_13 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_5, 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; uint8_t x_20; @@ -12085,7 +12085,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5645_() { +static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5615_() { _start: { lean_object* x_1; @@ -12201,7 +12201,7 @@ lean_dec(x_1); return x_2; } } -static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5701_() { +static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5641_() { _start: { lean_object* x_1; @@ -12218,7 +12218,7 @@ x_6 = l_Lean_registerBuiltinAttribute(x_5, x_4); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1() { _start: { lean_object* x_1; @@ -12226,24 +12226,24 @@ x_1 = lean_mk_string_from_bytes("invalid parser attribute implementation builder return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1; +x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_2) == 0) { lean_object* x_3; lean_dec(x_1); -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; return x_3; } else @@ -12262,7 +12262,7 @@ if (lean_obj_tag(x_5) == 0) lean_object* x_6; lean_dec(x_4); lean_dec(x_1); -x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; +x_6 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; return x_6; } else @@ -12297,7 +12297,7 @@ lean_dec(x_8); lean_dec(x_7); lean_dec(x_4); lean_dec(x_1); -x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; +x_13 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; return x_13; } } @@ -12308,7 +12308,7 @@ lean_dec(x_7); lean_dec(x_5); lean_dec(x_4); lean_dec(x_1); -x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; +x_14 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; return x_14; } } @@ -12319,13 +12319,13 @@ lean_object* x_15; lean_dec(x_4); lean_dec(x_2); lean_dec(x_1); -x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2; +x_15 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2; return x_15; } } } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1() { _start: { lean_object* x_1; @@ -12333,35 +12333,35 @@ x_1 = lean_mk_string_from_bytes("parserAttr", 10); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1; 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_Extension___hyg_5752____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1), 2, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1), 2, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3; x_4 = l_Lean_registerAttributeImplBuilder(x_2, x_3, x_1); return x_4; } } -static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5829_() { +static lean_object* _init_l___auto____x40_Lean_Parser_Extension___hyg_5739_() { _start: { lean_object* x_1; @@ -12397,7 +12397,7 @@ lean_ctor_set(x_14, 1, x_13); x_15 = lean_alloc_ctor(1, 2, 0); lean_ctor_set(x_15, 0, x_11); lean_ctor_set(x_15, 1, x_14); -x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2; +x_16 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2; x_17 = l_Lean_registerAttributeOfBuilder(x_9, x_16, x_5, x_15, x_10); return x_17; } @@ -12438,7 +12438,7 @@ x_8 = l_Lean_Parser_registerParserCategory(x_1, x_2, x_3, x_7, x_5, x_6); return x_8; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1() { _start: { lean_object* x_1; @@ -12446,17 +12446,17 @@ x_1 = lean_mk_string_from_bytes("builtin_term_parser", 19); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1; 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_Extension___hyg_5907____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3() { _start: { lean_object* x_1; @@ -12464,41 +12464,41 @@ x_1 = lean_mk_string_from_bytes("term", 4); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_mkParserOfConstantUnsafe___closed__4; x_2 = l_Lean_Parser_mkParserOfConstantUnsafe___closed__7; x_3 = l_Lean_Parser_registerBuiltinParserAttribute___closed__3; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__12; -x_2 = lean_unsigned_to_nat(5907u); +x_2 = lean_unsigned_to_nat(5787u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4; x_4 = 0; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5; x_6 = l_Lean_Parser_registerBuiltinParserAttribute(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1() { _start: { lean_object* x_1; @@ -12506,48 +12506,48 @@ x_1 = lean_mk_string_from_bytes("term_parser", 11); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1; 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_Extension___hyg_5946____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3; 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_Extension___hyg_5946____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__12; -x_2 = lean_unsigned_to_nat(5946u); +x_2 = lean_unsigned_to_nat(5826u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__3; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4; x_5 = l_Lean_Parser_registerBuiltinDynamicParserAttribute(x_2, x_3, x_4, x_1); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1() { _start: { lean_object* x_1; @@ -12555,17 +12555,17 @@ x_1 = lean_mk_string_from_bytes("builtin_command_parser", 22); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1; 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_Extension___hyg_5985____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3() { _start: { lean_object* x_1; @@ -12573,41 +12573,41 @@ x_1 = lean_mk_string_from_bytes("command", 7); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; x_1 = l_Lean_Parser_mkParserOfConstantUnsafe___closed__4; x_2 = l_Lean_Parser_mkParserOfConstantUnsafe___closed__7; x_3 = l_Lean_Parser_registerBuiltinParserAttribute___closed__3; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3; x_5 = l_Lean_Name_mkStr4(x_1, x_2, x_3, x_4); return x_5; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__12; -x_2 = lean_unsigned_to_nat(5985u); +x_2 = lean_unsigned_to_nat(5865u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; lean_object* x_5; lean_object* x_6; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4; x_4 = 0; -x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5; +x_5 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5; x_6 = l_Lean_Parser_registerBuiltinParserAttribute(x_2, x_3, x_4, x_5, x_1); return x_6; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1() { _start: { lean_object* x_1; @@ -12615,43 +12615,43 @@ x_1 = lean_mk_string_from_bytes("command_parser", 14); return x_1; } } -static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1; 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_Extension___hyg_6024____closed__3() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3; 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_Extension___hyg_6024____closed__4() { +static lean_object* _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_3698____closed__12; -x_2 = lean_unsigned_to_nat(6024u); +x_2 = lean_unsigned_to_nat(5904u); x_3 = l_Lean_Name_num___override(x_1, x_2); return x_3; } } -LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; lean_object* x_4; lean_object* x_5; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2; -x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3; -x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__4; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2; +x_3 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3; +x_4 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4; x_5 = l_Lean_Parser_registerBuiltinDynamicParserAttribute(x_2, x_3, x_4, x_1); return x_5; } @@ -12660,7 +12660,7 @@ LEAN_EXPORT lean_object* l_Lean_Parser_commandParser(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; -x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3; +x_2 = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3; x_3 = l_Lean_Parser_categoryParser(x_2, x_1); return x_3; } @@ -13490,7 +13490,7 @@ return x_8; } } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1() { _start: { lean_object* x_1; @@ -13498,21 +13498,21 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.ParserName.category", 31); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13520,7 +13520,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4() { _start: { lean_object* x_1; lean_object* x_2; @@ -13529,7 +13529,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5() { _start: { lean_object* x_1; lean_object* x_2; @@ -13538,7 +13538,7 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__6() { _start: { lean_object* x_1; @@ -13546,21 +13546,21 @@ x_1 = lean_mk_string_from_bytes("Lean.Parser.ParserName.parser", 29); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____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_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7; x_2 = lean_box(1); x_3 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_3, 0, x_1); @@ -13568,7 +13568,7 @@ lean_ctor_set(x_3, 1, x_2); return x_3; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9() { _start: { lean_object* x_1; @@ -13576,17 +13576,17 @@ x_1 = lean_mk_string_from_bytes("false", 5); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9; x_2 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11() { _start: { lean_object* x_1; @@ -13594,17 +13594,17 @@ x_1 = lean_mk_string_from_bytes("true", 4); return x_1; } } -static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12() { +static lean_object* _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11; +x_1 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11; 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_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489_(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369_(lean_object* x_1, lean_object* x_2) { _start: { if (lean_obj_tag(x_1) == 0) @@ -13616,14 +13616,14 @@ lean_dec(x_1); x_4 = lean_unsigned_to_nat(1024u); x_5 = lean_nat_dec_le(x_4, x_2); x_6 = l_Lean_Name_reprPrec(x_3, x_4); -x_7 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3; +x_7 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3; x_8 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_8, 0, x_7); lean_ctor_set(x_8, 1, x_6); if (x_5 == 0) { lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; -x_9 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4; +x_9 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4; x_10 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_10, 0, x_9); lean_ctor_set(x_10, 1, x_8); @@ -13637,7 +13637,7 @@ return x_13; else { lean_object* x_14; lean_object* x_15; uint8_t x_16; lean_object* x_17; lean_object* x_18; -x_14 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5; +x_14 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5; x_15 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_15, 0, x_14); lean_ctor_set(x_15, 1, x_8); @@ -13659,7 +13659,7 @@ lean_dec(x_1); x_21 = lean_unsigned_to_nat(1024u); x_22 = lean_nat_dec_le(x_21, x_2); x_23 = l_Lean_Name_reprPrec(x_19, x_21); -x_24 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8; +x_24 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8; x_25 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_25, 0, x_24); lean_ctor_set(x_25, 1, x_23); @@ -13672,11 +13672,11 @@ if (x_22 == 0) if (x_20 == 0) { 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; -x_28 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10; +x_28 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10; x_29 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_29, 0, x_27); lean_ctor_set(x_29, 1, x_28); -x_30 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4; +x_30 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4; x_31 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_31, 0, x_30); lean_ctor_set(x_31, 1, x_29); @@ -13690,11 +13690,11 @@ return x_34; else { 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; -x_35 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12; +x_35 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12; x_36 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_36, 0, x_27); lean_ctor_set(x_36, 1, x_35); -x_37 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4; +x_37 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4; x_38 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_38, 0, x_37); lean_ctor_set(x_38, 1, x_36); @@ -13711,11 +13711,11 @@ else if (x_20 == 0) { 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; -x_42 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10; +x_42 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10; x_43 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_43, 0, x_27); lean_ctor_set(x_43, 1, x_42); -x_44 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5; +x_44 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5; x_45 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_45, 0, x_44); lean_ctor_set(x_45, 1, x_43); @@ -13729,11 +13729,11 @@ return x_48; else { lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; uint8_t x_53; lean_object* x_54; lean_object* x_55; -x_49 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12; +x_49 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12; x_50 = lean_alloc_ctor(5, 2, 0); lean_ctor_set(x_50, 0, x_27); lean_ctor_set(x_50, 1, x_49); -x_51 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5; +x_51 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5; x_52 = lean_alloc_ctor(4, 2, 0); lean_ctor_set(x_52, 0, x_51); lean_ctor_set(x_52, 1, x_50); @@ -13748,11 +13748,11 @@ return x_55; } } } -LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____boxed(lean_object* x_1, lean_object* x_2) { +LEAN_EXPORT lean_object* l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____boxed(lean_object* x_1, lean_object* x_2) { _start: { lean_object* x_3; -x_3 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489_(x_1, x_2); +x_3 = l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369_(x_1, x_2); lean_dec(x_2); return x_3; } @@ -13761,7 +13761,7 @@ static lean_object* _init_l_Lean_Parser_instReprParserName___closed__1() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____boxed), 2, 0); +x_1 = lean_alloc_closure((void*)(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____boxed), 2, 0); return x_1; } } @@ -15475,75 +15475,75 @@ l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_ lean_mark_persistent(l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__3); l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__4 = _init_l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__4(); lean_mark_persistent(l_List_forM___at___private_Lean_Parser_Extension_0__Lean_Parser_ParserAttribute_add___spec__1___closed__4); -l___auto____x40_Lean_Parser_Extension___hyg_5645_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5645_(); -lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5645_); +l___auto____x40_Lean_Parser_Extension___hyg_5615_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5615_(); +lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5615_); l_Lean_Parser_mkParserAttributeImpl___closed__1 = _init_l_Lean_Parser_mkParserAttributeImpl___closed__1(); lean_mark_persistent(l_Lean_Parser_mkParserAttributeImpl___closed__1); l_Lean_Parser_mkParserAttributeImpl___closed__2 = _init_l_Lean_Parser_mkParserAttributeImpl___closed__2(); lean_mark_persistent(l_Lean_Parser_mkParserAttributeImpl___closed__2); -l___auto____x40_Lean_Parser_Extension___hyg_5701_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5701_(); -lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5701_); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____lambda__1___closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752____closed__3); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5752_(lean_io_mk_world()); +l___auto____x40_Lean_Parser_Extension___hyg_5641_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5641_(); +lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5641_); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____lambda__1___closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662____closed__3); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5662_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l___auto____x40_Lean_Parser_Extension___hyg_5829_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5829_(); -lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5829_); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907____closed__5); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5907_(lean_io_mk_world()); +}l___auto____x40_Lean_Parser_Extension___hyg_5739_ = _init_l___auto____x40_Lean_Parser_Extension___hyg_5739_(); +lean_mark_persistent(l___auto____x40_Lean_Parser_Extension___hyg_5739_); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787____closed__5); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5787_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946____closed__4); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5946_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826____closed__4); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5826_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__4); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985____closed__5); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5985_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__4); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865____closed__5); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5865_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); -}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__1); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__2); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__3); -l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__4(); -lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024____closed__4); -if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_6024_(lean_io_mk_world()); +}l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__1); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__2); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__3); +l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4 = _init_l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4(); +lean_mark_persistent(l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904____closed__4); +if (builtin) {res = l_Lean_Parser_initFn____x40_Lean_Parser_Extension___hyg_5904_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); }l_Lean_Parser_withOpenDeclFnCore___closed__1 = _init_l_Lean_Parser_withOpenDeclFnCore___closed__1(); @@ -15560,30 +15560,30 @@ l_Lean_Parser_withOpenFn___closed__1 = _init_l_Lean_Parser_withOpenFn___closed__ lean_mark_persistent(l_Lean_Parser_withOpenFn___closed__1); l_Lean_Parser_withOpenFn___closed__2 = _init_l_Lean_Parser_withOpenFn___closed__2(); lean_mark_persistent(l_Lean_Parser_withOpenFn___closed__2); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__1); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__2); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__3); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__4); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__5); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__6); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__7); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__8); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__9); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__10); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__11); -l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12(); -lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6489____closed__12); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__1); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__2); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__3); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__4); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__5); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__6 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__6(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__6); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__7); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__8); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__9); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__10); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__11); +l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12 = _init_l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12(); +lean_mark_persistent(l___private_Lean_Parser_Extension_0__Lean_Parser_reprParserName____x40_Lean_Parser_Extension___hyg_6369____closed__12); l_Lean_Parser_instReprParserName___closed__1 = _init_l_Lean_Parser_instReprParserName___closed__1(); lean_mark_persistent(l_Lean_Parser_instReprParserName___closed__1); l_Lean_Parser_instReprParserName = _init_l_Lean_Parser_instReprParserName(); diff --git a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c index bc2c662ae38a..aac6db616e83 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Formatter.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Formatter.c @@ -206,6 +206,7 @@ lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkLhsPrec_formatter___rarg(lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_hygieneInfoNoAntiquot_formatter___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instCoeForallForallFormatterAliasValue(lean_object*); uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Formatter_strLitNoAntiquot_formatter___closed__1; @@ -294,7 +295,6 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_instMonadTraverserFormat LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_visitArgs___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_lookahead_formatter___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_formatTactic___closed__1; -lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_nonReservedSymbolNoAntiquot_formatter(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Formatter_checkColGe_formatter(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_name_eq(lean_object*, lean_object*); @@ -11056,7 +11056,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_format(lean_object* x_1, lean_obje { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = l_Lean_PrettyPrinter_format___closed__2; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(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/PrettyPrinter/Parenthesizer.c b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c index 0b94fa972941..00007a6e6179 100644 --- a/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c +++ b/stage0/stdlib/Lean/PrettyPrinter/Parenthesizer.c @@ -206,6 +206,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_errorAtSavedPos_pare static lean_object* l_Lean_PrettyPrinter_parenthesizeCommand___closed__2; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkKind___closed__6; lean_object* lean_st_ref_take(lean_object*, lean_object*); +lean_object* l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1___closed__2; uint8_t lean_nat_dec_eq(lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_Parenthesizer_orelse_parenthesizer___closed__2; @@ -303,7 +304,6 @@ static lean_object* l_Lean_PrettyPrinter_Parenthesizer_maybeParenthesize___lambd static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___closed__1; static lean_object* l_Lean_PrettyPrinter_Parenthesizer_level_parenthesizer___lambda__1___closed__1; LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_withPosition_parenthesizer(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_rawIdentNoAntiquot_parenthesizer___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_Parenthesizer_checkColGt_parenthesizer___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_PrettyPrinter_mkCategoryParenthesizerAttribute___lambda__1___closed__1; @@ -10721,7 +10721,7 @@ LEAN_EXPORT lean_object* l_Lean_PrettyPrinter_parenthesize(lean_object* x_1, lea { lean_object* x_6; lean_object* x_7; lean_object* x_8; uint8_t x_9; x_6 = l_Lean_PrettyPrinter_parenthesize___closed__2; -x_7 = l_Lean_isTracingEnabledFor___at_Lean_addDecl___spec__8(x_6, x_3, x_4, x_5); +x_7 = l_Lean_isTracingEnabledFor___at_Lean_compileDecl___spec__3(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/ReducibilityAttrs.c b/stage0/stdlib/Lean/ReducibilityAttrs.c index 438ca3691fde..6761122fc3d5 100644 --- a/stage0/stdlib/Lean/ReducibilityAttrs.c +++ b/stage0/stdlib/Lean/ReducibilityAttrs.c @@ -20,6 +20,7 @@ static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_validate___closed static lean_object* l_Lean_getConstInfo___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_throwError___at___private_Lean_ReducibilityAttrs_0__Lean_validate___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_validate___closed__3; +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_allowUnsafeReducibility; static lean_object* l_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_326____closed__3; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_179____lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -40,6 +41,7 @@ uint8_t lean_usize_dec_le(size_t, size_t); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_326____spec__3___closed__3; static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_reprReducibilityStatus____x40_Lean_ReducibilityAttrs___hyg_15____closed__2; LEAN_EXPORT lean_object* l_Lean_instReprReducibilityStatus; +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute___boxed(lean_object*, lean_object*); size_t lean_uint64_to_usize(uint64_t); static lean_object* l_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_1404____closed__7; static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_validate___closed__10; @@ -222,6 +224,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_fold___at_Lean_initFn____x40_Lean_Reducib static lean_object* l___private_Lean_ReducibilityAttrs_0__Lean_validate___closed__9; LEAN_EXPORT lean_object* l_Lean_getReducibilityStatus___rarg___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* lean_set_reducibility_status(lean_object*, lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_mkHashMap___at_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_326____spec__12(lean_object*); static lean_object* l_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_326____closed__6; LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ReducibilityAttrs___hyg_1363____lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -5378,6 +5381,32 @@ x_4 = l_Lean_isIrreducible___rarg___lambda__1(x_1, x_3); return x_4; } } +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute___rarg(lean_object* x_1, lean_object* x_2) { +_start: +{ +uint8_t x_3; lean_object* x_4; +x_3 = 2; +x_4 = l_Lean_setReducibilityStatus___rarg(x_1, x_2, x_3); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = lean_alloc_closure((void*)(l_Lean_setIrreducibleAttribute___rarg), 2, 0); +return x_3; +} +} +LEAN_EXPORT lean_object* l_Lean_setIrreducibleAttribute___boxed(lean_object* x_1, lean_object* x_2) { +_start: +{ +lean_object* x_3; +x_3 = l_Lean_setIrreducibleAttribute(x_1, x_2); +lean_dec(x_2); +return x_3; +} +} lean_object* initialize_Lean_Attributes(uint8_t builtin, lean_object*); lean_object* initialize_Lean_ScopedEnvExtension(uint8_t builtin, lean_object*); static bool _G_initialized = false; diff --git a/stage0/stdlib/Lean/Replay.c b/stage0/stdlib/Lean/Replay.c index cf5071aa5129..5be46eb6e39b 100644 --- a/stage0/stdlib/Lean/Replay.c +++ b/stage0/stdlib/Lean/Replay.c @@ -1,6 +1,6 @@ // Lean compiler output // Module: Lean.Replay -// Imports: Lean.CoreM Lean.Util.FoldConsts +// Imports: Lean.CoreM Lean.AddDecl Lean.Util.FoldConsts #include #if defined(__clang__) #pragma clang diagnostic ignored "-Wunused-parameter" @@ -83,10 +83,9 @@ lean_object* lean_st_ref_get(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Environment_Replay_replayConstant___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_State_remaining___default; lean_object* lean_st_mk_ref(lean_object*, lean_object*); -lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__9; LEAN_EXPORT lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__2(lean_object*); +lean_object* l_Lean_throwKernelException___at_Lean_addDecl___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__18; lean_object* l_Lean_Name_str___override(lean_object*, lean_object*); uint8_t l___private_Lean_Declaration_0__Lean_beqConstructorVal____x40_Lean_Declaration___hyg_2031_(lean_object*, lean_object*); @@ -100,13 +99,12 @@ static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_check static lean_object* l_panic___at_Lean_Environment_Replay_replayConstant___spec__1___closed__4; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedRecursors(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedConstructors___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* lean_add_decl(lean_object*, lean_object*); +lean_object* l_Lean_Environment_addDecl(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_replayConstants___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__17; LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Environment_Replay_isTodo___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Environment_Replay_isTodo___spec__2___boxed(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Environment_Replay_replayConstant___spec__9(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_forIn_loop___at_Lean_Environment_Replay_replayConstant___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentHashMap_find_x3f___at_Lean_Environment_find_x3f___spec__2(lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__4; @@ -121,7 +119,6 @@ LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_Environment_Replay_replayCo LEAN_EXPORT lean_object* l_Lean_SMap_find_x3f___at_Lean_Environment_Replay_checkPostponedConstructors___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___boxed(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t l_Lean_NameSet_contains(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_panic_fn(lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__2; @@ -153,8 +150,6 @@ lean_object* lean_nat_add(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapM_loop___at_Lean_Environment_Replay_replayConstant___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___closed__2; uint8_t l_Lean_ConstantInfo_isPartial(lean_object*); -lean_object* l_Lean_KernelException_toMessageData(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Environment_Replay_throwKernelException___closed__10; LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Environment_Replay_checkPostponedRecursors___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Environment_Replay_checkPostponedConstructors(lean_object*, lean_object*, lean_object*); @@ -601,56 +596,6 @@ lean_dec(x_2); return x_5; } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__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; uint8_t x_7; -x_5 = lean_ctor_get(x_2, 5); -x_6 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_1, x_2, x_3, x_4); -x_7 = !lean_is_exclusive(x_6); -if (x_7 == 0) -{ -lean_object* x_8; lean_object* x_9; -x_8 = lean_ctor_get(x_6, 0); -lean_inc(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); -lean_ctor_set_tag(x_6, 1); -lean_ctor_set(x_6, 0, x_9); -return x_6; -} -else -{ -lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; -x_10 = lean_ctor_get(x_6, 0); -x_11 = lean_ctor_get(x_6, 1); -lean_inc(x_11); -lean_inc(x_10); -lean_dec(x_6); -lean_inc(x_5); -x_12 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_12, 0, x_5); -lean_ctor_set(x_12, 1, x_10); -x_13 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_13, 0, x_12); -lean_ctor_set(x_13, 1, x_11); -return x_13; -} -} -} -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; lean_object* x_6; lean_object* x_7; -x_5 = lean_ctor_get(x_2, 2); -lean_inc(x_5); -x_6 = l_Lean_KernelException_toMessageData(x_1, x_5); -x_7 = l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__2(x_6, x_2, x_3, x_4); -lean_dec(x_2); -return x_7; -} -} static lean_object* _init_l_Lean_Environment_Replay_throwKernelException___lambda__1___closed__1() { _start: { @@ -663,7 +608,6 @@ LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___lambda _start: { uint8_t x_8; -lean_dec(x_4); x_8 = !lean_is_exclusive(x_5); if (x_8 == 0) { @@ -677,7 +621,7 @@ x_12 = l_Lean_Option_get___at_Lean_profiler_threshold_getSecs___spec__1(x_1, x_1 lean_ctor_set(x_5, 4, x_12); lean_ctor_set(x_5, 2, x_1); lean_ctor_set_uint8(x_5, sizeof(void*)*11 + 1, x_2); -x_13 = l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1(x_3, x_5, x_6, x_7); +x_13 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_3, x_5, x_6, x_7); return x_13; } else @@ -719,7 +663,7 @@ lean_ctor_set(x_26, 9, x_21); lean_ctor_set(x_26, 10, x_22); lean_ctor_set_uint8(x_26, sizeof(void*)*11, x_23); lean_ctor_set_uint8(x_26, sizeof(void*)*11 + 1, x_2); -x_27 = l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1(x_3, x_26, x_6, x_7); +x_27 = l_Lean_throwKernelException___at_Lean_addDecl___spec__1(x_3, x_26, x_6, x_7); return x_27; } } @@ -1224,25 +1168,6 @@ goto block_26; } } } -LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_throwError___at_Lean_Environment_Replay_throwKernelException___spec__2(x_1, x_2, x_3, x_4); -lean_dec(x_3); -lean_dec(x_2); -return x_5; -} -} -LEAN_EXPORT lean_object* l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { -_start: -{ -lean_object* x_5; -x_5 = l_Lean_throwKernelException___at_Lean_Environment_Replay_throwKernelException___spec__1(x_1, x_2, x_3, x_4); -lean_dec(x_3); -return x_5; -} -} LEAN_EXPORT lean_object* l_Lean_Environment_Replay_throwKernelException___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: { @@ -1251,6 +1176,7 @@ x_8 = lean_unbox(x_2); lean_dec(x_2); x_9 = l_Lean_Environment_Replay_throwKernelException___lambda__1(x_1, x_8, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); +lean_dec(x_4); return x_9; } } @@ -1267,7 +1193,7 @@ return x_5; LEAN_EXPORT lean_object* l_Lean_Environment_Replay_addDecl(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { -lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; +lean_object* x_5; lean_object* x_6; lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; x_5 = lean_st_ref_get(x_3, x_4); x_6 = lean_ctor_get(x_5, 0); lean_inc(x_6); @@ -1277,97 +1203,98 @@ lean_dec(x_5); x_8 = lean_ctor_get(x_6, 0); lean_inc(x_8); lean_dec(x_6); -x_9 = lean_add_decl(x_8, x_1); -if (lean_obj_tag(x_9) == 0) +x_9 = lean_box(0); +x_10 = l_Lean_Environment_addDecl(x_8, x_9, x_1); +if (lean_obj_tag(x_10) == 0) { -lean_object* x_10; lean_object* x_11; -x_10 = lean_ctor_get(x_9, 0); -lean_inc(x_10); -lean_dec(x_9); -x_11 = l_Lean_Environment_Replay_throwKernelException(x_10, x_2, x_3, x_7); -return x_11; +lean_object* x_11; lean_object* x_12; +x_11 = lean_ctor_get(x_10, 0); +lean_inc(x_11); +lean_dec(x_10); +x_12 = l_Lean_Environment_Replay_throwKernelException(x_11, x_2, x_3, x_7); +return x_12; } else { -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_9, 0); -lean_inc(x_12); -lean_dec(x_9); -x_13 = lean_st_ref_take(x_3, x_7); -x_14 = lean_ctor_get(x_13, 0); -lean_inc(x_14); -x_15 = lean_ctor_get(x_13, 1); +lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; uint8_t x_17; +x_13 = lean_ctor_get(x_10, 0); +lean_inc(x_13); +lean_dec(x_10); +x_14 = lean_st_ref_take(x_3, x_7); +x_15 = lean_ctor_get(x_14, 0); lean_inc(x_15); -lean_dec(x_13); -x_16 = !lean_is_exclusive(x_14); -if (x_16 == 0) +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_17; lean_object* x_18; uint8_t x_19; -x_17 = lean_ctor_get(x_14, 0); -lean_dec(x_17); -lean_ctor_set(x_14, 0, x_12); -x_18 = lean_st_ref_set(x_3, x_14, x_15); -x_19 = !lean_is_exclusive(x_18); -if (x_19 == 0) +lean_object* x_18; lean_object* x_19; uint8_t x_20; +x_18 = lean_ctor_get(x_15, 0); +lean_dec(x_18); +lean_ctor_set(x_15, 0, x_13); +x_19 = lean_st_ref_set(x_3, x_15, x_16); +x_20 = !lean_is_exclusive(x_19); +if (x_20 == 0) { -lean_object* x_20; lean_object* x_21; -x_20 = lean_ctor_get(x_18, 0); -lean_dec(x_20); -x_21 = lean_box(0); -lean_ctor_set(x_18, 0, x_21); -return x_18; +lean_object* x_21; lean_object* x_22; +x_21 = lean_ctor_get(x_19, 0); +lean_dec(x_21); +x_22 = lean_box(0); +lean_ctor_set(x_19, 0, x_22); +return x_19; } else { -lean_object* x_22; lean_object* x_23; lean_object* x_24; -x_22 = lean_ctor_get(x_18, 1); -lean_inc(x_22); -lean_dec(x_18); -x_23 = lean_box(0); -x_24 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_24, 0, x_23); -lean_ctor_set(x_24, 1, x_22); -return x_24; +lean_object* x_23; lean_object* x_24; lean_object* x_25; +x_23 = lean_ctor_get(x_19, 1); +lean_inc(x_23); +lean_dec(x_19); +x_24 = lean_box(0); +x_25 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_23); +return x_25; } } else { -lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; -x_25 = lean_ctor_get(x_14, 1); -x_26 = lean_ctor_get(x_14, 2); -x_27 = lean_ctor_get(x_14, 3); -x_28 = lean_ctor_get(x_14, 4); +lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; 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_15, 1); +x_27 = lean_ctor_get(x_15, 2); +x_28 = lean_ctor_get(x_15, 3); +x_29 = lean_ctor_get(x_15, 4); +lean_inc(x_29); lean_inc(x_28); lean_inc(x_27); lean_inc(x_26); -lean_inc(x_25); -lean_dec(x_14); -x_29 = lean_alloc_ctor(0, 5, 0); -lean_ctor_set(x_29, 0, x_12); -lean_ctor_set(x_29, 1, x_25); -lean_ctor_set(x_29, 2, x_26); -lean_ctor_set(x_29, 3, x_27); -lean_ctor_set(x_29, 4, x_28); -x_30 = lean_st_ref_set(x_3, x_29, x_15); -x_31 = lean_ctor_get(x_30, 1); -lean_inc(x_31); -if (lean_is_exclusive(x_30)) { - lean_ctor_release(x_30, 0); - lean_ctor_release(x_30, 1); - x_32 = x_30; +lean_dec(x_15); +x_30 = lean_alloc_ctor(0, 5, 0); +lean_ctor_set(x_30, 0, x_13); +lean_ctor_set(x_30, 1, x_26); +lean_ctor_set(x_30, 2, x_27); +lean_ctor_set(x_30, 3, x_28); +lean_ctor_set(x_30, 4, x_29); +x_31 = lean_st_ref_set(x_3, x_30, x_16); +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_30); - x_32 = lean_box(0); + lean_dec_ref(x_31); + x_33 = lean_box(0); } -x_33 = lean_box(0); -if (lean_is_scalar(x_32)) { - x_34 = lean_alloc_ctor(0, 2, 0); +x_34 = lean_box(0); +if (lean_is_scalar(x_33)) { + x_35 = lean_alloc_ctor(0, 2, 0); } else { - x_34 = x_32; + x_35 = x_33; } -lean_ctor_set(x_34, 0, x_33); -lean_ctor_set(x_34, 1, x_31); -return x_34; +lean_ctor_set(x_35, 0, x_34); +lean_ctor_set(x_35, 1, x_32); +return x_35; } } } @@ -2120,7 +2047,7 @@ static lean_object* _init_l_Lean_Environment_Replay_replayConstant___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_Environment_Replay_replayConstant___closed__1; x_2 = l_Lean_Environment_Replay_replayConstant___closed__2; -x_3 = lean_unsigned_to_nat(75u); +x_3 = lean_unsigned_to_nat(76u); x_4 = lean_unsigned_to_nat(54u); x_5 = l_Lean_Environment_Replay_replayConstant___closed__3; x_6 = l___private_Init_Util_0__mkPanicMessageWithDecl(x_1, x_2, x_3, x_4, x_5); @@ -4895,6 +4822,7 @@ return x_7; } } lean_object* initialize_Lean_CoreM(uint8_t builtin, lean_object*); +lean_object* initialize_Lean_AddDecl(uint8_t builtin, lean_object*); lean_object* initialize_Lean_Util_FoldConsts(uint8_t builtin, lean_object*); static bool _G_initialized = false; LEAN_EXPORT lean_object* initialize_Lean_Replay(uint8_t builtin, lean_object* w) { @@ -4904,6 +4832,9 @@ _G_initialized = true; res = initialize_Lean_CoreM(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); +res = initialize_Lean_AddDecl(builtin, lean_io_mk_world()); +if (lean_io_result_is_error(res)) return res; +lean_dec_ref(res); res = initialize_Lean_Util_FoldConsts(builtin, lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); diff --git a/stage0/stdlib/Lean/ReservedNameAction.c b/stage0/stdlib/Lean/ReservedNameAction.c index f2fd3785260d..28b1e7f5633c 100644 --- a/stage0/stdlib/Lean/ReservedNameAction.c +++ b/stage0/stdlib/Lean/ReservedNameAction.c @@ -17,19 +17,24 @@ lean_object* l_Lean_Expr_const___override(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerReservedNameAction(lean_object*, lean_object*); lean_object* lean_format_pretty(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_mk_empty_array_with_capacity(lean_object*); +static lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; static lean_object* l_Lean_executeReservedNameAction___closed__1; static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_realizeGlobalConstNoOverload___spec__1___closed__1; static lean_object* l_Array_forInUnsafe_loop___at_Lean_executeReservedNameAction___spec__1___closed__2; static lean_object* l_Lean_ensureNonAmbiguous___at_Lean_realizeGlobalConstNoOverload___spec__1___closed__3; -lean_object* l_Lean_log___at_Lean_addDecl___spec__5(lean_object*, uint8_t, 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_filterFieldList___at_Lean_realizeGlobalConstCore___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__1; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_executeReservedNameAction___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); lean_object* lean_array_push(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_realizeGlobalConst(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_realizeGlobalConstNoOverload___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); static lean_object* l_panic___at_Lean_realizeGlobalConstNoOverload___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_stringToMessageData(lean_object*); @@ -41,28 +46,34 @@ static lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalCon size_t lean_usize_of_nat(lean_object*); static lean_object* l_Lean_registerReservedNameAction___closed__1; lean_object* l_List_mapTR_loop___at_Lean_filterFieldList___spec__2(lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_take(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_realizeGlobalConstNoOverload(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_executeReservedNameAction(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_realizeGlobalName___spec__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_realizeGlobalName___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_List_mapTR_loop___at_Lean_realizeGlobalConstNoOverloadCore___spec__4(lean_object*, lean_object*, lean_object*); lean_object* lean_st_ref_get(lean_object*, lean_object*); static lean_object* l_Lean_registerReservedNameAction___lambda__1___closed__1; uint8_t l_List_isEmpty___rarg(lean_object*); +LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_realizeGlobalName___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_ensureNonAmbiguous___at_Lean_realizeGlobalConstNoOverload___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_realizeGlobalConst___closed__1; LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstNoOverloadCore___spec__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*); -static lean_object* l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3; LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_executeReservedNameAction___spec__1(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_ReservedNameAction_0__Lean_reservedNameActionsRef; +extern lean_object* l_Lean_warningAsError; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2; lean_object* l___private_Init_Util_0__mkPanicMessageWithDecl(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerReservedNameAction___lambda__1___boxed(lean_object*, lean_object*, lean_object*); extern lean_object* l_Std_Format_defWidth; +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___boxed(lean_object*, 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_Lean_throwErrorAt___at_Lean_realizeGlobalConstNoOverload___spec__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_MessageData_ofExpr(lean_object*); @@ -99,12 +110,13 @@ lean_object* lean_array_get_size(lean_object*); LEAN_EXPORT lean_object* l_Lean_throwErrorAt___at_Lean_realizeGlobalConst___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__1; LEAN_EXPORT lean_object* l_Lean_resolveGlobalName___at_Lean_realizeGlobalName___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); uint8_t lean_usize_dec_lt(size_t, size_t); lean_object* lean_expr_dbg_to_string(lean_object*); LEAN_EXPORT lean_object* l_Lean_realizeGlobalName(lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_realizeGlobalName___spec__2(lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__2; uint8_t l_Lean_Exception_isRuntime(lean_object*); -LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_realizeGlobalName___spec__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstNoOverloadCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_registerReservedNameAction___lambda__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_filterFieldList___at_Lean_realizeGlobalConstCore___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); @@ -575,7 +587,926 @@ return x_18; } } } -LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_realizeGlobalName___spec__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +static lean_object* _init_l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_warningAsError; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(lean_object* x_1, lean_object* x_2, uint8_t x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +lean_object* x_7; uint8_t x_265; uint8_t x_266; +x_265 = 2; +x_266 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_265); +if (x_266 == 0) +{ +lean_object* x_267; +x_267 = lean_box(0); +x_7 = x_267; +goto block_264; +} +else +{ +lean_object* x_268; uint8_t x_269; +lean_inc(x_2); +x_268 = l_Lean_MessageData_hasSyntheticSorry(x_2); +x_269 = lean_unbox(x_268); +lean_dec(x_268); +if (x_269 == 0) +{ +lean_object* x_270; +x_270 = lean_box(0); +x_7 = x_270; +goto block_264; +} +else +{ +lean_object* x_271; lean_object* x_272; +lean_dec(x_2); +x_271 = lean_box(0); +x_272 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_6); +return x_272; +} +} +block_264: +{ +uint8_t x_8; lean_object* x_258; uint8_t x_259; uint8_t x_260; +lean_dec(x_7); +x_258 = lean_ctor_get(x_4, 2); +x_259 = 1; +x_260 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_259); +if (x_260 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +lean_object* x_261; uint8_t x_262; +x_261 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2; +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); +if (x_262 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +uint8_t x_263; +x_263 = 2; +x_8 = x_263; +goto block_257; +} +} +block_257: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_ctor_get(x_4, 5); +x_12 = lean_ctor_get(x_4, 6); +x_13 = lean_ctor_get(x_4, 7); +x_14 = l_Lean_replaceRef(x_1, x_11); +x_15 = 0; +x_16 = l_Lean_Syntax_getPos_x3f(x_14, x_15); +x_17 = l_Lean_Syntax_getTailPos_x3f(x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_18 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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_unsigned_to_nat(0u); +x_22 = l_Lean_FileMap_toPosition(x_10, x_21); +lean_inc(x_22); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_inc(x_13); +lean_inc(x_12); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_13); +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +x_26 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_27 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 3, x_26); +lean_ctor_set(x_27, 4, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_8); +x_28 = lean_st_ref_take(x_5, x_20); +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; uint8_t x_35; +x_32 = lean_ctor_get(x_29, 5); +x_33 = l_Lean_PersistentArray_push___rarg(x_32, x_27); +lean_ctor_set(x_29, 5, x_33); +x_34 = lean_st_ref_set(x_5, x_29, x_30); +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_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_29, 0); +x_42 = lean_ctor_get(x_29, 1); +x_43 = lean_ctor_get(x_29, 2); +x_44 = lean_ctor_get(x_29, 3); +x_45 = lean_ctor_get(x_29, 4); +x_46 = lean_ctor_get(x_29, 5); +x_47 = lean_ctor_get(x_29, 6); +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_inc(x_41); +lean_dec(x_29); +x_48 = l_Lean_PersistentArray_push___rarg(x_46, x_27); +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_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_48); +lean_ctor_set(x_49, 6, x_47); +x_50 = lean_st_ref_set(x_5, x_49, x_30); +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 +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_17); +if (x_55 == 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; lean_object* x_69; uint8_t x_70; +x_56 = lean_ctor_get(x_17, 0); +x_57 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l_Lean_FileMap_toPosition(x_10, x_60); +x_62 = l_Lean_FileMap_toPosition(x_10, x_56); +lean_dec(x_56); +lean_ctor_set(x_17, 0, x_62); +lean_inc(x_13); +lean_inc(x_12); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_12); +lean_ctor_set(x_63, 1, x_13); +x_64 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +x_65 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_66 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_66, 0, x_9); +lean_ctor_set(x_66, 1, x_61); +lean_ctor_set(x_66, 2, x_17); +lean_ctor_set(x_66, 3, x_65); +lean_ctor_set(x_66, 4, x_64); +lean_ctor_set_uint8(x_66, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_66, sizeof(void*)*5 + 1, x_8); +x_67 = lean_st_ref_take(x_5, x_59); +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_is_exclusive(x_68); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_68, 5); +x_72 = l_Lean_PersistentArray_push___rarg(x_71, x_66); +lean_ctor_set(x_68, 5, x_72); +x_73 = lean_st_ref_set(x_5, x_68, x_69); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +x_76 = lean_box(0); +lean_ctor_set(x_73, 0, x_76); +return x_73; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_80 = lean_ctor_get(x_68, 0); +x_81 = lean_ctor_get(x_68, 1); +x_82 = lean_ctor_get(x_68, 2); +x_83 = lean_ctor_get(x_68, 3); +x_84 = lean_ctor_get(x_68, 4); +x_85 = lean_ctor_get(x_68, 5); +x_86 = lean_ctor_get(x_68, 6); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_68); +x_87 = l_Lean_PersistentArray_push___rarg(x_85, x_66); +x_88 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 1, x_81); +lean_ctor_set(x_88, 2, x_82); +lean_ctor_set(x_88, 3, x_83); +lean_ctor_set(x_88, 4, x_84); +lean_ctor_set(x_88, 5, x_87); +lean_ctor_set(x_88, 6, x_86); +x_89 = lean_st_ref_set(x_5, x_88, x_69); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_91 = x_89; +} else { + lean_dec_ref(x_89); + x_91 = lean_box(0); +} +x_92 = lean_box(0); +if (lean_is_scalar(x_91)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_91; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +return x_93; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_94 = lean_ctor_get(x_17, 0); +lean_inc(x_94); +lean_dec(x_17); +x_95 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l_Lean_FileMap_toPosition(x_10, x_98); +x_100 = l_Lean_FileMap_toPosition(x_10, x_94); +lean_dec(x_94); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +lean_inc(x_13); +lean_inc(x_12); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_12); +lean_ctor_set(x_102, 1, x_13); +x_103 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_96); +x_104 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_105 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_105, 0, x_9); +lean_ctor_set(x_105, 1, x_99); +lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 3, x_104); +lean_ctor_set(x_105, 4, x_103); +lean_ctor_set_uint8(x_105, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_105, sizeof(void*)*5 + 1, x_8); +x_106 = lean_st_ref_take(x_5, x_97); +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_ctor_get(x_107, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +x_112 = lean_ctor_get(x_107, 3); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 4); +lean_inc(x_113); +x_114 = lean_ctor_get(x_107, 5); +lean_inc(x_114); +x_115 = lean_ctor_get(x_107, 6); +lean_inc(x_115); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + lean_ctor_release(x_107, 3); + lean_ctor_release(x_107, 4); + lean_ctor_release(x_107, 5); + lean_ctor_release(x_107, 6); + x_116 = x_107; +} else { + lean_dec_ref(x_107); + x_116 = lean_box(0); +} +x_117 = l_Lean_PersistentArray_push___rarg(x_114, x_105); +if (lean_is_scalar(x_116)) { + x_118 = lean_alloc_ctor(0, 7, 0); +} else { + x_118 = x_116; +} +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_110); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set(x_118, 3, x_112); +lean_ctor_set(x_118, 4, x_113); +lean_ctor_set(x_118, 5, x_117); +lean_ctor_set(x_118, 6, x_115); +x_119 = lean_st_ref_set(x_5, x_118, x_108); +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 +{ +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_124; +x_124 = !lean_is_exclusive(x_16); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_125 = lean_ctor_get(x_16, 0); +x_126 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_FileMap_toPosition(x_10, x_125); +lean_dec(x_125); +lean_inc(x_129); +lean_ctor_set(x_16, 0, x_129); +lean_inc(x_13); +lean_inc(x_12); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_12); +lean_ctor_set(x_130, 1, x_13); +x_131 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_127); +x_132 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_133 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_133, 0, x_9); +lean_ctor_set(x_133, 1, x_129); +lean_ctor_set(x_133, 2, x_16); +lean_ctor_set(x_133, 3, x_132); +lean_ctor_set(x_133, 4, x_131); +lean_ctor_set_uint8(x_133, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_133, sizeof(void*)*5 + 1, x_8); +x_134 = lean_st_ref_take(x_5, x_128); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = !lean_is_exclusive(x_135); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_138 = lean_ctor_get(x_135, 5); +x_139 = l_Lean_PersistentArray_push___rarg(x_138, x_133); +lean_ctor_set(x_135, 5, x_139); +x_140 = lean_st_ref_set(x_5, x_135, x_136); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_140, 0); +lean_dec(x_142); +x_143 = lean_box(0); +lean_ctor_set(x_140, 0, x_143); +return x_140; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_140, 1); +lean_inc(x_144); +lean_dec(x_140); +x_145 = lean_box(0); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_147 = lean_ctor_get(x_135, 0); +x_148 = lean_ctor_get(x_135, 1); +x_149 = lean_ctor_get(x_135, 2); +x_150 = lean_ctor_get(x_135, 3); +x_151 = lean_ctor_get(x_135, 4); +x_152 = lean_ctor_get(x_135, 5); +x_153 = lean_ctor_get(x_135, 6); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_135); +x_154 = l_Lean_PersistentArray_push___rarg(x_152, x_133); +x_155 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_155, 0, x_147); +lean_ctor_set(x_155, 1, x_148); +lean_ctor_set(x_155, 2, x_149); +lean_ctor_set(x_155, 3, x_150); +lean_ctor_set(x_155, 4, x_151); +lean_ctor_set(x_155, 5, x_154); +lean_ctor_set(x_155, 6, x_153); +x_156 = lean_st_ref_set(x_5, x_155, x_136); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_158 = x_156; +} else { + lean_dec_ref(x_156); + x_158 = lean_box(0); +} +x_159 = lean_box(0); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_158; +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_161 = lean_ctor_get(x_16, 0); +lean_inc(x_161); +lean_dec(x_16); +x_162 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_FileMap_toPosition(x_10, x_161); +lean_dec(x_161); +lean_inc(x_165); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +lean_inc(x_13); +lean_inc(x_12); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_12); +lean_ctor_set(x_167, 1, x_13); +x_168 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_163); +x_169 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_170 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_170, 0, x_9); +lean_ctor_set(x_170, 1, x_165); +lean_ctor_set(x_170, 2, x_166); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_168); +lean_ctor_set_uint8(x_170, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_170, sizeof(void*)*5 + 1, x_8); +x_171 = lean_st_ref_take(x_5, x_164); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_172, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_172, 2); +lean_inc(x_176); +x_177 = lean_ctor_get(x_172, 3); +lean_inc(x_177); +x_178 = lean_ctor_get(x_172, 4); +lean_inc(x_178); +x_179 = lean_ctor_get(x_172, 5); +lean_inc(x_179); +x_180 = lean_ctor_get(x_172, 6); +lean_inc(x_180); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + lean_ctor_release(x_172, 2); + lean_ctor_release(x_172, 3); + lean_ctor_release(x_172, 4); + lean_ctor_release(x_172, 5); + lean_ctor_release(x_172, 6); + x_181 = x_172; +} else { + lean_dec_ref(x_172); + x_181 = lean_box(0); +} +x_182 = l_Lean_PersistentArray_push___rarg(x_179, x_170); +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 7, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_174); +lean_ctor_set(x_183, 1, x_175); +lean_ctor_set(x_183, 2, x_176); +lean_ctor_set(x_183, 3, x_177); +lean_ctor_set(x_183, 4, x_178); +lean_ctor_set(x_183, 5, x_182); +lean_ctor_set(x_183, 6, x_180); +x_184 = lean_st_ref_set(x_5, x_183, x_173); +x_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = lean_box(0); +if (lean_is_scalar(x_186)) { + x_188 = lean_alloc_ctor(0, 2, 0); +} else { + x_188 = x_186; +} +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_185); +return x_188; +} +} +else +{ +lean_object* x_189; uint8_t x_190; +x_189 = lean_ctor_get(x_16, 0); +lean_inc(x_189); +lean_dec(x_16); +x_190 = !lean_is_exclusive(x_17); +if (x_190 == 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; 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; uint8_t x_204; +x_191 = lean_ctor_get(x_17, 0); +x_192 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_196 = l_Lean_FileMap_toPosition(x_10, x_191); +lean_dec(x_191); +lean_ctor_set(x_17, 0, x_196); +lean_inc(x_13); +lean_inc(x_12); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_12); +lean_ctor_set(x_197, 1, x_13); +x_198 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_193); +x_199 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_200 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_200, 0, x_9); +lean_ctor_set(x_200, 1, x_195); +lean_ctor_set(x_200, 2, x_17); +lean_ctor_set(x_200, 3, x_199); +lean_ctor_set(x_200, 4, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_200, sizeof(void*)*5 + 1, x_8); +x_201 = lean_st_ref_take(x_5, x_194); +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_is_exclusive(x_202); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; +x_205 = lean_ctor_get(x_202, 5); +x_206 = l_Lean_PersistentArray_push___rarg(x_205, x_200); +lean_ctor_set(x_202, 5, x_206); +x_207 = lean_st_ref_set(x_5, x_202, x_203); +x_208 = !lean_is_exclusive(x_207); +if (x_208 == 0) +{ +lean_object* x_209; lean_object* x_210; +x_209 = lean_ctor_get(x_207, 0); +lean_dec(x_209); +x_210 = lean_box(0); +lean_ctor_set(x_207, 0, x_210); +return x_207; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +lean_dec(x_207); +x_212 = lean_box(0); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_211); +return x_213; +} +} +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; +x_214 = lean_ctor_get(x_202, 0); +x_215 = lean_ctor_get(x_202, 1); +x_216 = lean_ctor_get(x_202, 2); +x_217 = lean_ctor_get(x_202, 3); +x_218 = lean_ctor_get(x_202, 4); +x_219 = lean_ctor_get(x_202, 5); +x_220 = lean_ctor_get(x_202, 6); +lean_inc(x_220); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_202); +x_221 = l_Lean_PersistentArray_push___rarg(x_219, x_200); +x_222 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_222, 0, x_214); +lean_ctor_set(x_222, 1, x_215); +lean_ctor_set(x_222, 2, x_216); +lean_ctor_set(x_222, 3, x_217); +lean_ctor_set(x_222, 4, x_218); +lean_ctor_set(x_222, 5, x_221); +lean_ctor_set(x_222, 6, x_220); +x_223 = lean_st_ref_set(x_5, x_222, x_203); +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_223); + x_225 = lean_box(0); +} +x_226 = lean_box(0); +if (lean_is_scalar(x_225)) { + x_227 = lean_alloc_ctor(0, 2, 0); +} else { + x_227 = x_225; +} +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_224); +return x_227; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_228 = lean_ctor_get(x_17, 0); +lean_inc(x_228); +lean_dec(x_17); +x_229 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_233 = l_Lean_FileMap_toPosition(x_10, x_228); +lean_dec(x_228); +x_234 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_234, 0, x_233); +lean_inc(x_13); +lean_inc(x_12); +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_12); +lean_ctor_set(x_235, 1, x_13); +x_236 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_236, 0, x_235); +lean_ctor_set(x_236, 1, x_230); +x_237 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; +lean_inc(x_9); +x_238 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_238, 0, x_9); +lean_ctor_set(x_238, 1, x_232); +lean_ctor_set(x_238, 2, x_234); +lean_ctor_set(x_238, 3, x_237); +lean_ctor_set(x_238, 4, x_236); +lean_ctor_set_uint8(x_238, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_238, sizeof(void*)*5 + 1, x_8); +x_239 = lean_st_ref_take(x_5, x_231); +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 = lean_ctor_get(x_240, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_240, 1); +lean_inc(x_243); +x_244 = lean_ctor_get(x_240, 2); +lean_inc(x_244); +x_245 = lean_ctor_get(x_240, 3); +lean_inc(x_245); +x_246 = lean_ctor_get(x_240, 4); +lean_inc(x_246); +x_247 = lean_ctor_get(x_240, 5); +lean_inc(x_247); +x_248 = lean_ctor_get(x_240, 6); +lean_inc(x_248); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + lean_ctor_release(x_240, 2); + lean_ctor_release(x_240, 3); + lean_ctor_release(x_240, 4); + lean_ctor_release(x_240, 5); + lean_ctor_release(x_240, 6); + x_249 = x_240; +} else { + lean_dec_ref(x_240); + x_249 = lean_box(0); +} +x_250 = l_Lean_PersistentArray_push___rarg(x_247, x_238); +if (lean_is_scalar(x_249)) { + x_251 = lean_alloc_ctor(0, 7, 0); +} else { + x_251 = x_249; +} +lean_ctor_set(x_251, 0, x_242); +lean_ctor_set(x_251, 1, x_243); +lean_ctor_set(x_251, 2, x_244); +lean_ctor_set(x_251, 3, x_245); +lean_ctor_set(x_251, 4, x_246); +lean_ctor_set(x_251, 5, x_250); +lean_ctor_set(x_251, 6, x_248); +x_252 = lean_st_ref_set(x_5, x_251, x_241); +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_254 = x_252; +} else { + lean_dec_ref(x_252); + x_254 = lean_box(0); +} +x_255 = lean_box(0); +if (lean_is_scalar(x_254)) { + x_256 = lean_alloc_ctor(0, 2, 0); +} else { + x_256 = x_254; +} +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_253); +return x_256; +} +} +} +} +} +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_realizeGlobalName___spec__2(lean_object* x_1, uint8_t x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +_start: +{ +lean_object* x_6; lean_object* x_7; +x_6 = lean_ctor_get(x_3, 5); +lean_inc(x_6); +x_7 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_6, x_1, x_2, x_3, x_4, x_5); +lean_dec(x_4); +lean_dec(x_3); +lean_dec(x_6); +return x_7; +} +} +LEAN_EXPORT lean_object* l_List_filterAuxM___at_Lean_realizeGlobalName___spec__4(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { if (lean_obj_tag(x_1) == 0) @@ -662,7 +1593,7 @@ x_33 = l_Lean_Exception_toMessageData(x_30); x_34 = 2; lean_inc(x_4); lean_inc(x_3); -x_35 = l_Lean_log___at_Lean_addDecl___spec__5(x_33, x_34, x_3, x_4, x_31); +x_35 = l_Lean_log___at_Lean_realizeGlobalName___spec__2(x_33, x_34, x_3, x_4, x_31); x_36 = lean_ctor_get(x_35, 1); lean_inc(x_36); lean_dec(x_35); @@ -693,7 +1624,7 @@ x_39 = l_Lean_Exception_toMessageData(x_30); x_40 = 2; lean_inc(x_4); lean_inc(x_3); -x_41 = l_Lean_log___at_Lean_addDecl___spec__5(x_39, x_40, x_3, x_4, x_31); +x_41 = l_Lean_log___at_Lean_realizeGlobalName___spec__2(x_39, x_40, x_3, x_4, x_31); x_42 = lean_ctor_get(x_41, 1); lean_inc(x_42); lean_dec(x_41); @@ -720,7 +1651,7 @@ x_47 = l_Lean_Exception_toMessageData(x_44); x_48 = 2; lean_inc(x_4); lean_inc(x_3); -x_49 = l_Lean_log___at_Lean_addDecl___spec__5(x_47, x_48, x_3, x_4, x_45); +x_49 = l_Lean_log___at_Lean_realizeGlobalName___spec__2(x_47, x_48, x_3, x_4, x_45); x_50 = lean_ctor_get(x_49, 1); lean_inc(x_50); lean_dec(x_49); @@ -754,7 +1685,7 @@ x_54 = l_Lean_Exception_toMessageData(x_44); x_55 = 2; lean_inc(x_4); lean_inc(x_3); -x_56 = l_Lean_log___at_Lean_addDecl___spec__5(x_54, x_55, x_3, x_4, x_45); +x_56 = l_Lean_log___at_Lean_realizeGlobalName___spec__2(x_54, x_55, x_3, x_4, x_45); x_57 = lean_ctor_get(x_56, 1); lean_inc(x_57); lean_dec(x_56); @@ -817,7 +1748,7 @@ x_7 = lean_ctor_get(x_5, 1); lean_inc(x_7); lean_dec(x_5); x_8 = lean_box(0); -x_9 = l_List_filterAuxM___at_Lean_realizeGlobalName___spec__2(x_6, x_8, x_2, x_3, x_7); +x_9 = l_List_filterAuxM___at_Lean_realizeGlobalName___spec__4(x_6, x_8, x_2, x_3, x_7); if (lean_obj_tag(x_9) == 0) { uint8_t x_10; @@ -878,6 +1809,29 @@ lean_dec(x_3); return x_5; } } +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +_start: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_8; +} +} +LEAN_EXPORT lean_object* l_Lean_log___at_Lean_realizeGlobalName___spec__2___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_6; lean_object* x_7; +x_6 = lean_unbox(x_2); +lean_dec(x_2); +x_7 = l_Lean_log___at_Lean_realizeGlobalName___spec__2(x_1, x_6, x_3, x_4, x_5); +return x_7; +} +} LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -1258,14 +2212,6 @@ x_1 = lean_mk_string_from_bytes("', possible interpretations: ", 29); return x_1; } } -static lean_object* _init_l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3() { -_start: -{ -lean_object* x_1; -x_1 = lean_mk_string_from_bytes("", 0); -return x_1; -} -} LEAN_EXPORT lean_object* l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { _start: { @@ -1286,7 +2232,7 @@ x_14 = l_List_toString___at_Lean_ensureNoOverload___spec__2(x_13); lean_dec(x_13); x_15 = lean_string_append(x_12, x_14); lean_dec(x_14); -x_16 = l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3; +x_16 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; x_17 = lean_string_append(x_15, x_16); x_18 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_18, 0, x_17); @@ -1334,7 +2280,7 @@ x_32 = l_List_toString___at_Lean_ensureNoOverload___spec__2(x_31); lean_dec(x_31); x_33 = lean_string_append(x_30, x_32); lean_dec(x_32); -x_34 = l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3; +x_34 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; x_35 = lean_string_append(x_33, x_34); x_36 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_36, 0, x_35); @@ -1900,7 +2846,7 @@ x_23 = l_List_toString___at_Lean_ensureNoOverload___spec__2(x_22); lean_dec(x_22); x_24 = lean_string_append(x_20, x_23); lean_dec(x_23); -x_25 = l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3; +x_25 = l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1; x_26 = lean_string_append(x_24, x_25); x_27 = lean_alloc_ctor(3, 1, 0); lean_ctor_set(x_27, 0, x_26); @@ -1998,6 +2944,10 @@ l_Array_forInUnsafe_loop___at_Lean_executeReservedNameAction___spec__1___closed_ lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_executeReservedNameAction___spec__1___closed__2); l_Lean_executeReservedNameAction___closed__1 = _init_l_Lean_executeReservedNameAction___closed__1(); lean_mark_persistent(l_Lean_executeReservedNameAction___closed__1); +l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1 = _init_l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1(); +lean_mark_persistent(l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__1); +l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2 = _init_l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2(); +lean_mark_persistent(l_Lean_logAt___at_Lean_realizeGlobalName___spec__3___closed__2); l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__1 = _init_l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__1(); lean_mark_persistent(l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__1); l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__2 = _init_l_Lean_throwUnknownConstant___at_Lean_realizeGlobalConstCore___spec__2___closed__2(); @@ -2010,8 +2960,6 @@ l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___c lean_mark_persistent(l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__1); l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__2 = _init_l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__2(); lean_mark_persistent(l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__2); -l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3 = _init_l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3(); -lean_mark_persistent(l_Lean_ensureNoOverload___at_Lean_realizeGlobalConstNoOverloadCore___spec__1___closed__3); l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__1 = _init_l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__1(); lean_mark_persistent(l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__1); l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__2 = _init_l_Lean_preprocessSyntaxAndResolve___at_Lean_realizeGlobalConst___spec__1___closed__2(); diff --git a/stage0/stdlib/Lean/ScopedEnvExtension.c b/stage0/stdlib/Lean/ScopedEnvExtension.c index 7184f53f626e..0dc3d1a989de 100644 --- a/stage0/stdlib/Lean/ScopedEnvExtension.c +++ b/stage0/stdlib/Lean/ScopedEnvExtension.c @@ -133,7 +133,6 @@ lean_object* l_panic___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insert___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__12(lean_object*); static lean_object* l_Lean_PersistentHashMap_insertAux___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__13___rarg___closed__1; static lean_object* l_Lean_registerScopedEnvExtensionUnsafe___rarg___closed__2; -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1038_(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_ScopedEnvExtension_activateScoped___spec__10___rarg(lean_object*, lean_object*); lean_object* l_Array_reverse___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_addScopedEntry___rarg(lean_object*, lean_object*, lean_object*, lean_object*); @@ -160,6 +159,7 @@ LEAN_EXPORT lean_object* l_Lean_SimpleScopedEnvExtension_Descr_finalizeImport___ lean_object* l_Lean_PersistentEnvExtension_addEntry___rarg(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_PersistentEnvExtension_getState___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAtCollisionNodeAux___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__26(lean_object*); +LEAN_EXPORT lean_object* l___auto____x40_Lean_ScopedEnvExtension___hyg_2269_; LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_insertAux_traverse___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__14(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_activateScoped___spec__1(lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_find_x3f___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__8___rarg(lean_object*, lean_object*); @@ -328,6 +328,7 @@ lean_object* lean_array_uget(lean_object*, size_t); LEAN_EXPORT lean_object* l_Lean_HashMap_insert___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__16___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_AssocList_contains___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__17___rarg___boxed(lean_object*, lean_object*); lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1008_(lean_object*); size_t lean_usize_shift_left(size_t, size_t); lean_object* l_Lean_Name_mkStr4(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_HashMapImp_expand___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__29___rarg(lean_object*, lean_object*); @@ -387,7 +388,6 @@ lean_object* l_List_toArrayAux___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_ScopedEnvExtension_activateScoped___spec__15___rarg(lean_object*, lean_object*, lean_object*, size_t, size_t, lean_object*); lean_object* l___private_Lean_Data_HashMap_0__Lean_numBucketsForCapacity(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_ScopedEnvExtension_addImportedFn___spec__3___rarg___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l___auto____x40_Lean_ScopedEnvExtension___hyg_2299_; LEAN_EXPORT lean_object* l_Lean_ScopedEnvExtension_instInhabitedStateStack(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_PersistentHashMap_findAtAux___at_Lean_ScopedEnvExtension_ScopedEntries_insert___spec__4(lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_popScope___spec__1___rarg___lambda__1(lean_object*, size_t, lean_object*, lean_object*, lean_object*, size_t, lean_object*); @@ -5243,7 +5243,7 @@ lean_dec(x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1038_(lean_object* x_1) { +LEAN_EXPORT lean_object* l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1008_(lean_object* x_1) { _start: { lean_object* x_2; lean_object* x_3; uint8_t x_4; @@ -8087,7 +8087,7 @@ x_11 = l_Array_forInUnsafe_loop___at_Lean_activateScoped___spec__1___rarg(x_1, x return x_11; } } -static lean_object* _init_l___auto____x40_Lean_ScopedEnvExtension___hyg_2299_() { +static lean_object* _init_l___auto____x40_Lean_ScopedEnvExtension___hyg_2269_() { _start: { lean_object* x_1; @@ -8303,7 +8303,7 @@ l_Lean_instInhabitedScopedEnvExtension___rarg___closed__5 = _init_l_Lean_instInh lean_mark_persistent(l_Lean_instInhabitedScopedEnvExtension___rarg___closed__5); l_Lean_instInhabitedScopedEnvExtension___rarg___closed__6 = _init_l_Lean_instInhabitedScopedEnvExtension___rarg___closed__6(); lean_mark_persistent(l_Lean_instInhabitedScopedEnvExtension___rarg___closed__6); -if (builtin) {res = l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1038_(lean_io_mk_world()); +if (builtin) {res = l_Lean_initFn____x40_Lean_ScopedEnvExtension___hyg_1008_(lean_io_mk_world()); if (lean_io_result_is_error(res)) return res; l_Lean_scopedEnvExtensionsRef = lean_io_result_get_value(res); lean_mark_persistent(l_Lean_scopedEnvExtensionsRef); @@ -8332,8 +8332,8 @@ l_Array_forInUnsafe_loop___at_Lean_pushScope___spec__1___rarg___lambda__1___clos lean_mark_persistent(l_Array_forInUnsafe_loop___at_Lean_pushScope___spec__1___rarg___lambda__1___closed__1); l_Lean_pushScope___rarg___closed__1 = _init_l_Lean_pushScope___rarg___closed__1(); lean_mark_persistent(l_Lean_pushScope___rarg___closed__1); -l___auto____x40_Lean_ScopedEnvExtension___hyg_2299_ = _init_l___auto____x40_Lean_ScopedEnvExtension___hyg_2299_(); -lean_mark_persistent(l___auto____x40_Lean_ScopedEnvExtension___hyg_2299_); +l___auto____x40_Lean_ScopedEnvExtension___hyg_2269_ = _init_l___auto____x40_Lean_ScopedEnvExtension___hyg_2269_(); +lean_mark_persistent(l___auto____x40_Lean_ScopedEnvExtension___hyg_2269_); l_Lean_registerSimpleScopedEnvExtension___rarg___closed__1 = _init_l_Lean_registerSimpleScopedEnvExtension___rarg___closed__1(); lean_mark_persistent(l_Lean_registerSimpleScopedEnvExtension___rarg___closed__1); return lean_io_result_mk_ok(lean_box(0)); diff --git a/stage0/stdlib/Lean/Server/CodeActions/Attr.c b/stage0/stdlib/Lean/Server/CodeActions/Attr.c index fd176fb7df60..96b9c150cfc3 100644 --- a/stage0/stdlib/Lean/Server/CodeActions/Attr.c +++ b/stage0/stdlib/Lean/Server/CodeActions/Attr.c @@ -88,6 +88,7 @@ LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_CodeAction_mkHoleCodeAction___s lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_82____lambda__4___boxed(lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Attr_0__Lean_CodeAction_addBuiltin___closed__4; +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); size_t lean_usize_of_nat(lean_object*); static lean_object* l___private_Lean_Server_CodeActions_Attr_0__Lean_CodeAction_addBuiltin___closed__27; static lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_1052____lambda__1___closed__1; @@ -208,7 +209,6 @@ LEAN_EXPORT lean_object* l_Lean_CodeAction_CommandCodeActions_insert(lean_object lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_793____lambda__3(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_297____closed__3; -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_CodeAction_mkCommandCodeAction_unsafe__1___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_297____closed__14; LEAN_EXPORT lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions_Attr___hyg_1052____lambda__3(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); @@ -1163,7 +1163,7 @@ x_8 = lean_ctor_get(x_7, 1); lean_inc(x_8); lean_dec(x_7); x_9 = 0; -x_10 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_9); +x_10 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_9); if (x_10 == 0) { lean_object* x_11; lean_object* x_12; uint8_t x_13; @@ -3024,7 +3024,7 @@ LEAN_EXPORT lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions { uint8_t x_7; uint8_t x_8; x_7 = 0; -x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_7); +x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_7); if (x_8 == 0) { lean_object* x_9; lean_object* x_10; uint8_t x_11; @@ -3982,7 +3982,7 @@ LEAN_EXPORT lean_object* l_Lean_CodeAction_initFn____x40_Lean_Server_CodeActions { uint8_t x_7; uint8_t x_8; x_7 = 0; -x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_3, x_7); +x_8 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_3, x_7); if (x_8 == 0) { lean_object* x_9; lean_object* x_10; uint8_t x_11; diff --git a/stage0/stdlib/Lean/Server/CodeActions/Basic.c b/stage0/stdlib/Lean/Server/CodeActions/Basic.c index 7b30bb4a87c6..5468ba158a49 100644 --- a/stage0/stdlib/Lean/Server/CodeActions/Basic.c +++ b/stage0/stdlib/Lean/Server/CodeActions/Basic.c @@ -112,6 +112,7 @@ lean_object* l_Lean_initializing(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_CodeAction_getFileSource_x21___boxed(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_1509____spec__1___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Name_mkStr3(lean_object*, lean_object*, lean_object*); +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_instFileSourceCodeAction(lean_object*); static lean_object* l_Lean_Server_CodeAction_getFileSource_x21___closed__4; size_t lean_usize_of_nat(lean_object*); @@ -228,7 +229,6 @@ static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___ LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_1509____spec__1___lambda__1(lean_object*); lean_object* l_Lean_getConstInfo___at_Lean_registerInitAttrUnsafe___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_parseRequestParams___at_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_1179____spec__2(lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_registerLspRequestHandler___at_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_1179____spec__1___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_initFn____x40_Lean_Server_CodeActions_Basic___hyg_494____closed__1; @@ -1526,7 +1526,7 @@ x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); x_10 = 0; -x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_10); +x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; @@ -1759,7 +1759,7 @@ x_9 = lean_ctor_get(x_8, 1); lean_inc(x_9); lean_dec(x_8); x_10 = 0; -x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_10); +x_11 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(x_4, x_10); if (x_11 == 0) { lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; uint8_t x_18; diff --git a/stage0/stdlib/Lean/Server/FileWorker.c b/stage0/stdlib/Lean/Server/FileWorker.c index e511d3fe09a9..54d7c97b615a 100644 --- a/stage0/stdlib/Lean/Server/FileWorker.c +++ b/stage0/stdlib/Lean/Server/FileWorker.c @@ -20,7 +20,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_updateDocument(lean_object*, l lean_object* l_Lean_Language_Lean_process(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_ImportCompletion_collectAvailableImports(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___boxed(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_getObjValAs_x3f___at___private_Lean_Message_0__Lean_fromJsonBaseMessage____x40_Lean_Message___hyg_2259____spec__1(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_288____closed__5; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__3; @@ -49,6 +49,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleResponse(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_setupImports___closed__2; lean_object* l_Lean_JsonNumber_toString(lean_object*); +lean_object* l_List_forIn_loop___at_Lean_KVMap_mergeBy___spec__1(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_server_reportDelayMs; LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__43; @@ -56,9 +57,9 @@ LEAN_EXPORT lean_object* l_IO_FS_Stream_readNotificationAs___at_Lean_Server_File extern lean_object* l_Lean_Server_FileWorker_RpcSession_keepAliveTimeMs; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__1(lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Server_documentUriFromModule(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4(lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRpcConnect___boxed(lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__6; static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__12; LEAN_EXPORT lean_object* l_Lean_RBNode_del___at_Lean_Server_FileWorker_mainLoop___spec__3___boxed(lean_object*, lean_object*); @@ -71,7 +72,7 @@ lean_object* l_Lean_PersistentArray_toArray___rarg(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_initAndRunWorker_writeError___closed__1; lean_object* lean_io_check_canceled(lean_object*); -lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_Language_mkIncrementalProcessor___rarg(lean_object*, lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_forInUnsafe_loop___at_Lean_Server_FileWorker_handleRpcRelease___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* lean_io_has_finished(lean_object*, lean_object*); @@ -80,6 +81,7 @@ LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_handleRpc LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_handleStaleDependency___spec__1(lean_object*, lean_object*, lean_object*); lean_object* lean_io_promise_new(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_sendServerRequest(lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__51; static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__13; static lean_object* l_Lean_Server_FileWorker_initFn____x40_Lean_Server_FileWorker___hyg_288____closed__4; @@ -135,7 +137,6 @@ lean_object* l_Lean_Option_register___at_Lean_initFn____x40_Lean_Util_RecDepth__ lean_object* l_Lean_Name_mkStr5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__2; lean_object* l___private_Lean_Widget_InteractiveDiagnostic_0__Lean_Widget_Lean_Lsp_DiagnosticWith_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveDiagnostic___hyg_2223_(lean_object*); -static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__53; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__6; lean_object* l_Lean_Json_opt___at_Lean_JsonRpc_instToJsonMessage___spec__1(lean_object*, lean_object*); @@ -231,8 +232,10 @@ lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_toJsonPublishDiag LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_ReportSnapshotsState_allInfoTrees___default; static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__5; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleStaleDependency(lean_object*); +static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3; static lean_object* l_Lean_Server_FileWorker_parseParams___rarg___closed__2; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__5___boxed(lean_object*); +static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1; lean_object* l_IO_Mutex_atomically___at_IO_Channel_recvAllCurrent___spec__1___rarg(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__7; lean_object* l_Lean_RBNode_setBlack___rarg(lean_object*); @@ -278,7 +281,6 @@ lean_object* l_Array_mapMUnsafe_map___at_Lean_Widget_Lean_Lsp_DiagnosticWith_ins static lean_object* l_Lean_Server_FileWorker_handleRequest___closed__5; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__15; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots___lambda__5(lean_object*, lean_object*); -static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3; static lean_object* l_Lean_Server_FileWorker_initializeWorker___closed__1; lean_object* l_Lean_FileMap_ofString(lean_object*); lean_object* lean_st_mk_ref(lean_object*, lean_object*); @@ -324,6 +326,7 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__67; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__22; LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots___lambda__6___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__1___closed__5; static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__8; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleRequest___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -363,7 +366,7 @@ static lean_object* l_Lean_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___s LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots_go___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__1; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_handleRpcConnect___spec__1___boxed(lean_object*, lean_object*, lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_mkIleanInfoFinalNotification___closed__1; static lean_object* l_Lean_Server_FileWorker_initAndRunWorker___closed__1; LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at_Lean_Server_FileWorker_handleRequest___spec__5(size_t, size_t, lean_object*, lean_object*); @@ -381,6 +384,7 @@ lean_object* l_IO_FS_Stream_writeLspMessage(lean_object*, lean_object*, lean_obj lean_object* l___private_Lean_Data_Lsp_Diagnostics_0__Lean_Lsp_fromJsonPublishDiagnosticsParams____x40_Lean_Data_Lsp_Diagnostics___hyg_2430_(lean_object*); LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanDidOpenTextDocumentParams____x40_Lean_Data_Lsp_Extra___hyg_187_(lean_object*); +lean_object* l_Lean_searchModuleNameOfFileName(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_runRefreshTask___closed__1; LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___rarg(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2(lean_object*, lean_object*, lean_object*, lean_object*); @@ -401,7 +405,7 @@ LEAN_EXPORT lean_object* l___private_Lean_Server_FileWorker_0__Lean_Server_FileW static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__26; lean_object* l_Lean_Name_mkStr2(lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_handleRequest___closed__2; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__3___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Language_Lean_instToSnapshotTreeCommandParsedSnapshot_go(lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__31; lean_object* l_Lean_Environment_allImportedModuleNames(lean_object*); @@ -434,7 +438,7 @@ uint8_t l_Lean_RBNode_isBlack___rarg(lean_object*); lean_object* lean_io_bind_task(lean_object*, lean_object*, lean_object*, uint8_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_instInhabitedReportSnapshotsState; -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Array_filterMapM___at_Lean_Server_FileWorker_initializeWorker___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_parseParams___at_Lean_Server_FileWorker_handleNotification___spec__4(lean_object*, lean_object*, lean_object*, lean_object*); @@ -448,7 +452,9 @@ static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__1___closed__ static lean_object* l_Lean_Widget_Lean_Lsp_DiagnosticWith_instRpcEncodableDiagnosticWith_enc____x40_Lean_Widget_InteractiveDiagnostic___hyg_1549____at_Lean_Server_FileWorker_handleRequest___spec__4___closed__9; uint8_t lean_uint64_dec_eq(uint64_t, uint64_t); lean_object* l_IO_Channel_new___rarg(lean_object*); +static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__29; +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6(lean_object*, lean_object*, lean_object*); static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__28; LEAN_EXPORT lean_object* l_Lean_RBNode_insert___at_Lean_Server_FileWorker_handleRpcConnect___spec__1(lean_object*, uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_foldM___at_Lean_Server_FileWorker_mainLoop___spec__1___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); @@ -466,7 +472,6 @@ static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_ini static lean_object* l_Lean_Server_FileWorker_handleRequest___closed__4; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__50; lean_object* lean_array_uget(lean_object*, size_t); -static lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2; lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonRpcConnectParams____x40_Lean_Data_Lsp_Extra___hyg_1629_(lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleRequest___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_IO_FS_Stream_withPrefix(lean_object*, lean_object*); @@ -495,7 +500,7 @@ LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_workerMain___boxed__const__1; LEAN_EXPORT lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2(lean_object*); static lean_object* l_Lean_Json_toStructured_x3f___at_Lean_Server_FileWorker_runRefreshTask___spec__2___closed__2; lean_object* lean_array_get_size(lean_object*); -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_erase___at_Lean_Server_FileWorker_mainLoop___spec__2(uint64_t, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_forIn_visit___at_Lean_Server_FileWorker_mainLoop___spec__4(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_handleGetInteractiveDiagnosticsRequest(lean_object*, lean_object*, lean_object*, lean_object*); @@ -508,6 +513,7 @@ static lean_object* l_Lean_Loop_forIn_loop___at_Lean_Server_FileWorker_runRefres lean_object* lean_int_neg(lean_object*); LEAN_EXPORT lean_object* l_Array_foldlMUnsafe_fold___at_Lean_Server_FileWorker_handleGetInteractiveDiagnosticsRequest___spec__1___boxed(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_Server_FileWorker_setupImports___lambda__6___closed__1; static lean_object* l_IO_FS_Stream_readRequestAs___at_Lean_Server_FileWorker_initAndRunWorker___spec__2___closed__71; uint8_t lean_usize_dec_lt(size_t, size_t); static lean_object* l_Lean_Server_FileWorker_handleNotification___closed__2; @@ -527,7 +533,6 @@ lean_object* l_Lean_Json_pretty(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_RBNode_ins___at_Lean_Server_FileWorker_handleRpcConnect___spec__2___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_FileWorker_setupImports___closed__1; lean_object* l___private_Lean_Widget_TaggedText_0__Lean_Widget_toJsonTaggedText____x40_Lean_Widget_TaggedText___hyg_617____at_Lean_Widget_TaggedText_instRpcEncodable___spec__3(lean_object*); -static lean_object* l_Lean_Server_FileWorker_initializeWorker___closed__2; LEAN_EXPORT lean_object* l_IO_ofExcept___at_Lean_Server_FileWorker_handleRequest___spec__3(lean_object*, lean_object*); lean_object* l___private_Lean_Data_Lsp_Extra_0__Lean_Lsp_fromJsonLeanFileProgressParams____x40_Lean_Data_Lsp_Extra___hyg_827_(lean_object*); lean_object* l___private_Lean_Data_Lsp_Internal_0__Lean_Lsp_toJsonLeanIleanInfoParams____x40_Lean_Data_Lsp_Internal___hyg_1782_(lean_object*); @@ -3296,46 +3301,126 @@ return x_29; } } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; lean_object* x_7; uint8_t x_8; -x_6 = lean_ctor_get(x_1, 1); -lean_inc(x_6); -x_7 = lean_io_promise_resolve(x_6, x_2, x_5); -x_8 = !lean_is_exclusive(x_7); -if (x_8 == 0) +lean_inc(x_3); +return x_3; +} +} +static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1() { +_start: { -lean_object* x_9; lean_object* x_10; lean_object* x_11; -x_9 = lean_ctor_get(x_7, 0); -lean_dec(x_9); -x_10 = lean_ctor_get(x_1, 2); -lean_inc(x_10); +lean_object* x_1; +x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports___lambda__2___boxed), 3, 0); +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___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; uint32_t x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +x_6 = lean_ctor_get(x_1, 2); +lean_inc(x_6); lean_dec(x_1); +x_7 = l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1; +x_8 = l_List_forIn_loop___at_Lean_KVMap_mergeBy___spec__1(x_7, x_6, x_2); +x_9 = 0; +x_10 = lean_alloc_ctor(0, 2, 4); +lean_ctor_set(x_10, 0, x_3); +lean_ctor_set(x_10, 1, x_8); +lean_ctor_set_uint32(x_10, sizeof(void*)*2, x_9); x_11 = lean_alloc_ctor(1, 1, 0); lean_ctor_set(x_11, 0, x_10); -lean_ctor_set(x_7, 0, x_11); -return x_7; +x_12 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_12, 0, x_11); +lean_ctor_set(x_12, 1, x_5); +return x_12; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___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) { +_start: +{ +lean_object* x_8; lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; +lean_dec(x_5); +x_8 = lean_ctor_get(x_1, 1); +lean_inc(x_8); +lean_inc(x_8); +x_9 = lean_io_promise_resolve(x_8, x_2, x_7); +lean_dec(x_2); +x_10 = lean_ctor_get(x_9, 1); +lean_inc(x_10); +lean_dec(x_9); +x_11 = lean_ctor_get(x_4, 0); +lean_inc(x_11); +lean_dec(x_4); +x_12 = l_System_Uri_fileUriToPath_x3f(x_11); +if (lean_obj_tag(x_12) == 0) +{ +lean_object* x_13; lean_object* x_14; +lean_dec(x_8); +x_13 = lean_box(0); +x_14 = l_Lean_Server_FileWorker_setupImports___lambda__3(x_1, x_3, x_13, x_6, x_10); +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_7, 1); -lean_inc(x_12); -lean_dec(x_7); -x_13 = lean_ctor_get(x_1, 2); -lean_inc(x_13); -lean_dec(x_1); -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_12); -return x_15; +lean_object* x_15; lean_object* x_16; lean_object* x_17; +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +lean_dec(x_12); +lean_inc(x_15); +x_16 = l_Lean_searchModuleNameOfFileName(x_15, x_8, x_10); +lean_dec(x_8); +x_17 = lean_ctor_get(x_16, 0); +lean_inc(x_17); +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_box(0); +x_20 = lean_module_name_of_file(x_15, x_19, x_18); +if (lean_obj_tag(x_20) == 0) +{ +lean_object* x_21; lean_object* x_22; lean_object* x_23; +x_21 = lean_ctor_get(x_20, 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_Server_FileWorker_setupImports___lambda__3(x_1, x_3, x_21, x_6, x_22); +return x_23; +} +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 = l_Lean_Server_FileWorker_setupImports___lambda__3(x_1, x_3, x_25, x_6, x_24); +return x_26; } } +else +{ +lean_object* x_27; lean_object* x_28; lean_object* x_29; +lean_dec(x_15); +x_27 = lean_ctor_get(x_16, 1); +lean_inc(x_27); +lean_dec(x_16); +x_28 = lean_ctor_get(x_17, 0); +lean_inc(x_28); +lean_dec(x_17); +x_29 = l_Lean_Server_FileWorker_setupImports___lambda__3(x_1, x_3, x_28, x_6, x_27); +return x_29; } -static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1() { +} +} +} +static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1() { _start: { lean_object* x_1; @@ -3343,195 +3428,208 @@ x_1 = lean_mk_string_from_bytes("Imports are out of date and must be rebuilt; us return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__3(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__5(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7, lean_object* x_8) { _start: { -lean_object* x_8; lean_object* x_9; lean_object* x_10; -lean_dec(x_5); -x_8 = l_Lean_Elab_headerToImports(x_1); +lean_object* x_9; lean_object* x_10; lean_object* x_11; +lean_dec(x_6); +x_9 = l_Lean_Elab_headerToImports(x_1); lean_dec(x_1); lean_inc(x_3); lean_inc(x_2); -x_9 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports___lambda__1), 4, 2); -lean_closure_set(x_9, 0, x_2); -lean_closure_set(x_9, 1, x_3); +x_10 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports___lambda__1), 4, 2); +lean_closure_set(x_10, 0, x_2); +lean_closure_set(x_10, 1, x_3); lean_inc(x_2); -x_10 = l_Lean_Server_FileWorker_setupFile(x_2, x_8, x_9, x_7); -if (lean_obj_tag(x_10) == 0) +x_11 = l_Lean_Server_FileWorker_setupFile(x_2, x_9, x_10, x_8); +if (lean_obj_tag(x_11) == 0) { -lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_52; 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_11 = lean_ctor_get(x_10, 0); -lean_inc(x_11); -x_12 = lean_ctor_get(x_10, 1); +lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_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_12 = lean_ctor_get(x_11, 0); lean_inc(x_12); -lean_dec(x_10); -x_52 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_ReportSnapshotsState_allInfoTrees___default___closed__1; -x_53 = l_Lean_Server_mkPublishDiagnosticsNotification(x_2, x_52); -x_54 = lean_ctor_get(x_53, 1); -lean_inc(x_54); -lean_dec(x_53); -x_55 = l_Lean_Json_toStructured_x3f___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishDiagnostics___spec__4(x_54); -x_56 = lean_ctor_get(x_55, 0); -lean_inc(x_56); -lean_dec(x_55); -x_57 = lean_alloc_ctor(1, 1, 0); -lean_ctor_set(x_57, 0, x_56); -x_58 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishDiagnostics___closed__1; -x_59 = lean_alloc_ctor(1, 2, 0); -lean_ctor_set(x_59, 0, x_58); -lean_ctor_set(x_59, 1, x_57); -x_60 = l_IO_Channel_send___rarg(x_3, x_59, x_12); -x_61 = lean_ctor_get(x_60, 1); -lean_inc(x_61); -lean_dec(x_60); -x_13 = x_61; -goto block_51; -block_51: +x_13 = lean_ctor_get(x_11, 1); +lean_inc(x_13); +lean_dec(x_11); +x_53 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_ReportSnapshotsState_allInfoTrees___default___closed__1; +lean_inc(x_2); +x_54 = l_Lean_Server_mkPublishDiagnosticsNotification(x_2, x_53); +x_55 = lean_ctor_get(x_54, 1); +lean_inc(x_55); +lean_dec(x_54); +x_56 = l_Lean_Json_toStructured_x3f___at___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishDiagnostics___spec__4(x_55); +x_57 = lean_ctor_get(x_56, 0); +lean_inc(x_57); +lean_dec(x_56); +x_58 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_58, 0, x_57); +x_59 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_publishDiagnostics___closed__1; +x_60 = lean_alloc_ctor(1, 2, 0); +lean_ctor_set(x_60, 0, x_59); +lean_ctor_set(x_60, 1, x_58); +x_61 = l_IO_Channel_send___rarg(x_3, x_60, x_13); +x_62 = lean_ctor_get(x_61, 1); +lean_inc(x_62); +lean_dec(x_61); +x_14 = x_62; +goto block_52; +block_52: { -lean_object* x_14; -x_14 = lean_ctor_get(x_11, 0); -lean_inc(x_14); -switch (lean_obj_tag(x_14)) { +lean_object* x_15; +x_15 = lean_ctor_get(x_12, 0); +lean_inc(x_15); +switch (lean_obj_tag(x_15)) { case 2: { -lean_object* x_15; lean_object* x_16; uint8_t x_17; -lean_dec(x_11); -x_15 = l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1; -x_16 = l_Lean_Language_diagnosticsOfHeaderError(x_15, x_6, x_13); -x_17 = !lean_is_exclusive(x_16); -if (x_17 == 0) +lean_object* x_16; lean_object* x_17; uint8_t x_18; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_16 = l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1; +x_17 = l_Lean_Language_diagnosticsOfHeaderError(x_16, x_7, x_14); +lean_dec(x_7); +x_18 = !lean_is_exclusive(x_17); +if (x_18 == 0) { -lean_object* x_18; lean_object* x_19; uint8_t x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; -x_18 = lean_ctor_get(x_16, 0); -x_19 = lean_box(0); -x_20 = 1; -x_21 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_21, 0, x_18); -lean_ctor_set(x_21, 1, x_19); -lean_ctor_set_uint8(x_21, sizeof(void*)*2, x_20); -x_22 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_22, 0, x_21); -lean_ctor_set(x_22, 1, x_19); -x_23 = lean_alloc_ctor(0, 1, 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; +x_19 = lean_ctor_get(x_17, 0); +x_20 = lean_box(0); +x_21 = 1; +x_22 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_22, 0, x_19); +lean_ctor_set(x_22, 1, x_20); +lean_ctor_set_uint8(x_22, sizeof(void*)*2, x_21); +x_23 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_23, 0, x_22); -lean_ctor_set(x_16, 0, x_23); -return x_16; +lean_ctor_set(x_23, 1, x_20); +x_24 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_24, 0, x_23); +lean_ctor_set(x_17, 0, x_24); +return x_17; } else { -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; -x_24 = lean_ctor_get(x_16, 0); -x_25 = lean_ctor_get(x_16, 1); +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; lean_object* x_32; +x_25 = lean_ctor_get(x_17, 0); +x_26 = lean_ctor_get(x_17, 1); +lean_inc(x_26); lean_inc(x_25); -lean_inc(x_24); -lean_dec(x_16); -x_26 = lean_box(0); -x_27 = 1; -x_28 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_28, 0, x_24); -lean_ctor_set(x_28, 1, x_26); -lean_ctor_set_uint8(x_28, sizeof(void*)*2, 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_26); -x_30 = lean_alloc_ctor(0, 1, 0); +lean_dec(x_17); +x_27 = lean_box(0); +x_28 = 1; +x_29 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_29, 0, x_25); +lean_ctor_set(x_29, 1, x_27); +lean_ctor_set_uint8(x_29, sizeof(void*)*2, x_28); +x_30 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_30, 0, x_29); -x_31 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_30, 1, x_27); +x_31 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_31, 0, x_30); -lean_ctor_set(x_31, 1, x_25); -return x_31; +x_32 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_32, 0, x_31); +lean_ctor_set(x_32, 1, x_26); +return x_32; } } case 3: { -lean_object* x_32; lean_object* x_33; uint8_t x_34; -lean_dec(x_11); -x_32 = lean_ctor_get(x_14, 0); -lean_inc(x_32); -lean_dec(x_14); -x_33 = l_Lean_Language_diagnosticsOfHeaderError(x_32, x_6, x_13); -x_34 = !lean_is_exclusive(x_33); -if (x_34 == 0) +lean_object* x_33; lean_object* x_34; uint8_t x_35; +lean_dec(x_12); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_2); +x_33 = lean_ctor_get(x_15, 0); +lean_inc(x_33); +lean_dec(x_15); +x_34 = l_Lean_Language_diagnosticsOfHeaderError(x_33, x_7, x_14); +lean_dec(x_7); +x_35 = !lean_is_exclusive(x_34); +if (x_35 == 0) { -lean_object* x_35; lean_object* x_36; uint8_t x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; -x_35 = lean_ctor_get(x_33, 0); -x_36 = lean_box(0); -x_37 = 1; -x_38 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_38, 0, x_35); -lean_ctor_set(x_38, 1, x_36); -lean_ctor_set_uint8(x_38, sizeof(void*)*2, x_37); -x_39 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_39, 0, x_38); -lean_ctor_set(x_39, 1, x_36); -x_40 = lean_alloc_ctor(0, 1, 0); +lean_object* x_36; lean_object* x_37; uint8_t x_38; lean_object* x_39; lean_object* x_40; lean_object* x_41; +x_36 = lean_ctor_get(x_34, 0); +x_37 = lean_box(0); +x_38 = 1; +x_39 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_39, 0, x_36); +lean_ctor_set(x_39, 1, x_37); +lean_ctor_set_uint8(x_39, sizeof(void*)*2, x_38); +x_40 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_40, 0, x_39); -lean_ctor_set(x_33, 0, x_40); -return x_33; +lean_ctor_set(x_40, 1, x_37); +x_41 = lean_alloc_ctor(0, 1, 0); +lean_ctor_set(x_41, 0, x_40); +lean_ctor_set(x_34, 0, x_41); +return x_34; } 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; lean_object* x_47; lean_object* x_48; -x_41 = lean_ctor_get(x_33, 0); -x_42 = lean_ctor_get(x_33, 1); +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; +x_42 = lean_ctor_get(x_34, 0); +x_43 = lean_ctor_get(x_34, 1); +lean_inc(x_43); lean_inc(x_42); -lean_inc(x_41); -lean_dec(x_33); -x_43 = lean_box(0); -x_44 = 1; -x_45 = lean_alloc_ctor(0, 2, 1); -lean_ctor_set(x_45, 0, x_41); -lean_ctor_set(x_45, 1, x_43); -lean_ctor_set_uint8(x_45, sizeof(void*)*2, x_44); -x_46 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_46, 0, x_45); -lean_ctor_set(x_46, 1, x_43); -x_47 = lean_alloc_ctor(0, 1, 0); +lean_dec(x_34); +x_44 = lean_box(0); +x_45 = 1; +x_46 = lean_alloc_ctor(0, 2, 1); +lean_ctor_set(x_46, 0, x_42); +lean_ctor_set(x_46, 1, x_44); +lean_ctor_set_uint8(x_46, sizeof(void*)*2, x_45); +x_47 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_47, 0, x_46); -x_48 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_47, 1, x_44); +x_48 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_48, 0, x_47); -lean_ctor_set(x_48, 1, x_42); -return x_48; +x_49 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_49, 0, x_48); +lean_ctor_set(x_49, 1, x_43); +return x_49; } } default: { -lean_object* x_49; lean_object* x_50; -lean_dec(x_14); -x_49 = lean_box(0); -x_50 = l_Lean_Server_FileWorker_setupImports___lambda__2(x_11, x_4, x_49, x_6, x_13); -return x_50; +lean_object* x_50; lean_object* x_51; +lean_dec(x_15); +x_50 = lean_box(0); +x_51 = l_Lean_Server_FileWorker_setupImports___lambda__4(x_12, x_4, x_5, x_2, x_50, x_7, x_14); +lean_dec(x_7); +return x_51; } } } } else { -uint8_t x_62; +uint8_t x_63; +lean_dec(x_7); +lean_dec(x_5); +lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); -x_62 = !lean_is_exclusive(x_10); -if (x_62 == 0) +x_63 = !lean_is_exclusive(x_11); +if (x_63 == 0) { -return x_10; +return x_11; } else { -lean_object* x_63; lean_object* x_64; lean_object* x_65; -x_63 = lean_ctor_get(x_10, 0); -x_64 = lean_ctor_get(x_10, 1); +lean_object* x_64; lean_object* x_65; lean_object* x_66; +x_64 = lean_ctor_get(x_11, 0); +x_65 = lean_ctor_get(x_11, 1); +lean_inc(x_65); lean_inc(x_64); -lean_inc(x_63); -lean_dec(x_10); -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_dec(x_11); +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; } } } } -static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1() { +static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__1() { _start: { lean_object* x_1; lean_object* x_2; uint8_t x_3; lean_object* x_4; @@ -3545,33 +3643,33 @@ lean_ctor_set_uint8(x_4, sizeof(void*)*2, x_3); return x_4; } } -static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2() { +static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2() { _start: { lean_object* x_1; lean_object* x_2; lean_object* x_3; x_1 = lean_box(0); -x_2 = l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1; +x_2 = l_Lean_Server_FileWorker_setupImports___lambda__6___closed__1; x_3 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_3, 0, x_2); lean_ctor_set(x_3, 1, x_1); return x_3; } } -static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3() { +static lean_object* _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3() { _start: { lean_object* x_1; lean_object* x_2; -x_1 = l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2; +x_1 = l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2; x_2 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_2, 0, x_1); return x_2; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; lean_object* x_5; -x_4 = l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3; +x_4 = l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3; x_5 = lean_alloc_ctor(0, 2, 0); lean_ctor_set(x_5, 0, x_4); lean_ctor_set(x_5, 1, x_3); @@ -3590,138 +3688,145 @@ static lean_object* _init_l_Lean_Server_FileWorker_setupImports___closed__2() { _start: { lean_object* x_1; -x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports___lambda__4___boxed), 3, 0); +x_1 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports___lambda__6___boxed), 3, 0); return x_1; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6, lean_object* x_7) { _start: { -lean_object* x_7; lean_object* x_8; lean_object* x_9; lean_object* x_10; uint8_t x_11; lean_object* x_12; lean_object* x_13; uint8_t x_14; -x_7 = l_Lean_Server_FileWorker_setupImports___closed__1; -x_8 = lean_st_ref_take(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_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; uint8_t x_15; +x_8 = l_Lean_Server_FileWorker_setupImports___closed__1; +x_9 = lean_st_ref_take(x_8, x_7); +x_10 = lean_ctor_get(x_9, 0); lean_inc(x_10); -lean_dec(x_8); -x_11 = 1; -x_12 = lean_box(x_11); -x_13 = lean_st_ref_set(x_7, x_12, x_10); -x_14 = lean_unbox(x_9); +x_11 = lean_ctor_get(x_9, 1); +lean_inc(x_11); lean_dec(x_9); -if (x_14 == 0) +x_12 = 1; +x_13 = lean_box(x_12); +x_14 = lean_st_ref_set(x_8, x_13, x_11); +x_15 = lean_unbox(x_10); +lean_dec(x_10); +if (x_15 == 0) { -lean_object* x_15; lean_object* x_16; lean_object* x_17; -x_15 = lean_ctor_get(x_13, 1); -lean_inc(x_15); -lean_dec(x_13); -x_16 = lean_box(0); -x_17 = l_Lean_Server_FileWorker_setupImports___lambda__3(x_4, x_1, x_2, x_3, x_16, x_5, x_15); -lean_dec(x_5); -lean_dec(x_3); -return x_17; +lean_object* x_16; lean_object* x_17; lean_object* x_18; +x_16 = lean_ctor_get(x_14, 1); +lean_inc(x_16); +lean_dec(x_14); +x_17 = lean_box(0); +x_18 = l_Lean_Server_FileWorker_setupImports___lambda__5(x_5, x_1, x_3, x_4, x_2, x_17, x_6, x_16); +return x_18; } else { -lean_object* x_18; uint32_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; uint8_t x_26; +lean_object* x_19; uint32_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; uint8_t x_27; +lean_dec(x_5); lean_dec(x_4); lean_dec(x_3); lean_dec(x_2); lean_dec(x_1); -x_18 = lean_ctor_get(x_13, 1); -lean_inc(x_18); -lean_dec(x_13); -x_19 = 200; -x_20 = l_IO_sleep(x_19, x_18); -x_21 = lean_ctor_get(x_20, 1); -lean_inc(x_21); -lean_dec(x_20); -x_22 = lean_io_check_canceled(x_21); -x_23 = lean_ctor_get(x_22, 0); -lean_inc(x_23); -x_24 = lean_ctor_get(x_22, 1); +x_19 = lean_ctor_get(x_14, 1); +lean_inc(x_19); +lean_dec(x_14); +x_20 = 200; +x_21 = l_IO_sleep(x_20, x_19); +x_22 = lean_ctor_get(x_21, 1); +lean_inc(x_22); +lean_dec(x_21); +x_23 = lean_io_check_canceled(x_22); +x_24 = lean_ctor_get(x_23, 0); lean_inc(x_24); -lean_dec(x_22); -x_25 = l_Lean_Server_FileWorker_setupImports___closed__2; -x_26 = lean_unbox(x_23); +x_25 = lean_ctor_get(x_23, 1); +lean_inc(x_25); lean_dec(x_23); -if (x_26 == 0) +x_26 = l_Lean_Server_FileWorker_setupImports___closed__2; +x_27 = lean_unbox(x_24); +lean_dec(x_24); +if (x_27 == 0) { -uint8_t x_27; lean_object* x_28; -x_27 = 2; -x_28 = lean_io_exit(x_27, x_24); -if (lean_obj_tag(x_28) == 0) +uint8_t x_28; lean_object* x_29; +x_28 = 2; +x_29 = lean_io_exit(x_28, x_25); +if (lean_obj_tag(x_29) == 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_object* x_30; lean_object* x_31; lean_object* x_32; +x_30 = lean_ctor_get(x_29, 0); lean_inc(x_30); -lean_dec(x_28); -x_31 = lean_apply_3(x_25, x_29, x_5, x_30); -return x_31; +x_31 = lean_ctor_get(x_29, 1); +lean_inc(x_31); +lean_dec(x_29); +x_32 = lean_apply_3(x_26, x_30, x_6, x_31); +return x_32; } else { -uint8_t x_32; -lean_dec(x_5); -x_32 = !lean_is_exclusive(x_28); -if (x_32 == 0) +uint8_t x_33; +lean_dec(x_6); +x_33 = !lean_is_exclusive(x_29); +if (x_33 == 0) { -return x_28; +return x_29; } else { -lean_object* x_33; lean_object* x_34; lean_object* x_35; -x_33 = lean_ctor_get(x_28, 0); -x_34 = lean_ctor_get(x_28, 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_28); -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_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 { -lean_object* x_36; lean_object* x_37; -x_36 = lean_box(0); -x_37 = lean_apply_3(x_25, x_36, x_5, x_24); -return x_37; +lean_object* x_37; lean_object* x_38; +x_37 = lean_box(0); +x_38 = lean_apply_3(x_26, x_37, x_6, x_25); +return x_38; } } } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__2___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { -lean_object* x_6; -x_6 = l_Lean_Server_FileWorker_setupImports___lambda__2(x_1, x_2, x_3, x_4, x_5); -lean_dec(x_4); +lean_object* x_4; +x_4 = l_Lean_Server_FileWorker_setupImports___lambda__2(x_1, x_2, x_3); lean_dec(x_3); lean_dec(x_2); +lean_dec(x_1); +return x_4; +} +} +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___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_FileWorker_setupImports___lambda__3(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_setupImports___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_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___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_Server_FileWorker_setupImports___lambda__3(x_1, x_2, x_3, x_4, x_5, x_6, x_7); +x_8 = l_Lean_Server_FileWorker_setupImports___lambda__4(x_1, x_2, x_3, x_4, x_5, x_6, x_7); lean_dec(x_6); -lean_dec(x_4); return x_8; } } -LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__4___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { +LEAN_EXPORT lean_object* l_Lean_Server_FileWorker_setupImports___lambda__6___boxed(lean_object* x_1, lean_object* x_2, lean_object* x_3) { _start: { lean_object* x_4; -x_4 = l_Lean_Server_FileWorker_setupImports___lambda__4(x_1, x_2, x_3); +x_4 = l_Lean_Server_FileWorker_setupImports___lambda__6(x_1, x_2, x_3); lean_dec(x_2); lean_dec(x_1); return x_4; @@ -4661,282 +4766,212 @@ x_2 = lean_nat_to_int(x_1); return x_2; } } -static lean_object* _init_l_Lean_Server_FileWorker_initializeWorker___closed__2() { -_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_Server_FileWorker_initializeWorker(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4, lean_object* x_5, lean_object* x_6) { _start: { lean_object* x_7; lean_object* x_8; lean_object* x_9; uint8_t x_10; x_7 = lean_ctor_get(x_4, 3); lean_inc(x_7); -x_8 = lean_ctor_get(x_1, 0); -lean_inc(x_8); -x_9 = l_System_Uri_fileUriToPath_x3f(x_8); +x_8 = l_Lean_Server_FileWorker_initializeWorker___closed__1; +x_9 = lean_st_mk_ref(x_8, x_6); if (lean_obj_tag(x_7) == 0) { -uint8_t x_87; -x_87 = 0; -x_10 = x_87; -goto block_86; +uint8_t x_69; +x_69 = 0; +x_10 = x_69; +goto block_68; } else { -lean_object* x_88; lean_object* x_89; -x_88 = lean_ctor_get(x_7, 0); -lean_inc(x_88); +lean_object* x_70; lean_object* x_71; +x_70 = lean_ctor_get(x_7, 0); +lean_inc(x_70); lean_dec(x_7); -x_89 = lean_ctor_get(x_88, 1); -lean_inc(x_89); -lean_dec(x_88); -if (lean_obj_tag(x_89) == 0) +x_71 = lean_ctor_get(x_70, 1); +lean_inc(x_71); +lean_dec(x_70); +if (lean_obj_tag(x_71) == 0) { -uint8_t x_90; -x_90 = 0; -x_10 = x_90; -goto block_86; +uint8_t x_72; +x_72 = 0; +x_10 = x_72; +goto block_68; } else { -lean_object* x_91; uint8_t x_92; -x_91 = lean_ctor_get(x_89, 0); -lean_inc(x_91); -lean_dec(x_89); -x_92 = lean_unbox(x_91); -lean_dec(x_91); -x_10 = x_92; -goto block_86; -} +lean_object* x_73; uint8_t x_74; +x_73 = lean_ctor_get(x_71, 0); +lean_inc(x_73); +lean_dec(x_71); +x_74 = lean_unbox(x_73); +lean_dec(x_73); +x_10 = x_74; +goto block_68; } -block_86: -{ -lean_object* x_11; lean_object* x_12; -if (lean_obj_tag(x_9) == 0) -{ -lean_object* x_76; -x_76 = l_Lean_Server_FileWorker_initializeWorker___closed__2; -x_11 = x_76; -x_12 = x_6; -goto block_75; } -else +block_68: { -lean_object* x_77; lean_object* x_78; lean_object* x_79; -x_77 = lean_ctor_get(x_9, 0); -lean_inc(x_77); +lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; lean_object* x_36; lean_object* x_37; lean_object* x_38; lean_object* x_39; lean_object* x_40; uint8_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; lean_object* x_52; lean_object* x_53; lean_object* x_54; uint8_t x_55; +x_11 = lean_ctor_get(x_9, 0); +lean_inc(x_11); +x_12 = lean_ctor_get(x_9, 1); +lean_inc(x_12); lean_dec(x_9); -x_78 = lean_box(0); -x_79 = lean_module_name_of_file(x_77, x_78, x_6); -if (lean_obj_tag(x_79) == 0) -{ -lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; -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_box(0); -x_83 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_83, 0, x_82); -lean_ctor_set(x_83, 1, x_80); -x_11 = x_83; -x_12 = x_81; -goto block_75; -} -else -{ -lean_object* x_84; lean_object* x_85; -x_84 = lean_ctor_get(x_79, 1); -lean_inc(x_84); -lean_dec(x_79); -x_85 = l_Lean_Server_FileWorker_initializeWorker___closed__2; -x_11 = x_85; -x_12 = x_84; -goto block_75; -} -} -block_75: -{ -lean_object* x_13; lean_object* x_14; lean_object* x_15; lean_object* x_16; lean_object* x_17; lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; lean_object* x_25; lean_object* x_26; lean_object* x_27; lean_object* x_28; lean_object* x_29; lean_object* x_30; lean_object* x_31; lean_object* x_32; lean_object* x_33; lean_object* x_34; lean_object* x_35; uint32_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; lean_object* x_45; lean_object* x_46; lean_object* x_47; uint8_t x_48; lean_object* x_49; lean_object* x_50; lean_object* x_51; lean_object* x_52; 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; -x_13 = lean_ctor_get(x_11, 1); -lean_inc(x_13); -lean_dec(x_11); -x_14 = l_Lean_Server_FileWorker_initializeWorker___closed__1; -x_15 = lean_st_mk_ref(x_14, x_12); -x_16 = lean_ctor_get(x_15, 0); -lean_inc(x_16); -x_17 = lean_ctor_get(x_15, 1); +x_13 = lean_st_mk_ref(x_8, 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_IO_Channel_new___rarg(x_15); +x_17 = lean_ctor_get(x_16, 0); lean_inc(x_17); -lean_dec(x_15); -x_18 = lean_st_mk_ref(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 = l_IO_Channel_new___rarg(x_20); -x_22 = lean_ctor_get(x_21, 0); +x_18 = lean_ctor_get(x_16, 1); +lean_inc(x_18); +lean_dec(x_16); +x_19 = lean_box(0); +x_20 = lean_st_mk_ref(x_19, 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); -x_23 = lean_ctor_get(x_21, 1); -lean_inc(x_23); -lean_dec(x_21); -x_24 = lean_box(0); -x_25 = lean_st_mk_ref(x_24, x_23); -x_26 = lean_ctor_get(x_25, 0); -lean_inc(x_26); -x_27 = lean_ctor_get(x_25, 1); +lean_dec(x_20); +lean_inc(x_17); +lean_inc(x_11); +x_23 = l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel(x_2, x_11, x_17, x_22); +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_io_promise_new(x_25); +x_27 = lean_ctor_get(x_26, 0); lean_inc(x_27); -lean_dec(x_25); -lean_inc(x_22); -lean_inc(x_16); -x_28 = l_Lean_Server_FileWorker_initializeWorker_mkLspOutputChannel(x_2, x_16, x_22, 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_io_promise_new(x_30); +x_28 = lean_ctor_get(x_26, 1); +lean_inc(x_28); +lean_dec(x_26); +lean_inc(x_27); +lean_inc(x_24); +lean_inc(x_5); +lean_inc(x_1); +x_29 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports), 7, 4); +lean_closure_set(x_29, 0, x_1); +lean_closure_set(x_29, 1, x_5); +lean_closure_set(x_29, 2, x_24); +lean_closure_set(x_29, 3, x_27); +x_30 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process), 4, 1); +lean_closure_set(x_30, 0, x_29); +x_31 = l_Lean_Language_mkIncrementalProcessor___rarg(x_30, x_28); 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); -lean_inc(x_32); -lean_inc(x_29); -lean_inc(x_1); -x_34 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_setupImports), 6, 3); -lean_closure_set(x_34, 0, x_1); -lean_closure_set(x_34, 1, x_29); -lean_closure_set(x_34, 2, x_32); -x_35 = lean_alloc_closure((void*)(l_Lean_Language_Lean_process), 4, 1); -lean_closure_set(x_35, 0, x_34); -x_36 = 0; -lean_inc(x_5); -x_37 = lean_alloc_ctor(0, 2, 4); -lean_ctor_set(x_37, 0, x_13); -lean_ctor_set(x_37, 1, x_5); -lean_ctor_set_uint32(x_37, sizeof(void*)*2, x_36); -x_38 = l_Lean_Language_mkIncrementalProcessor___rarg(x_35, x_37, x_33); -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); lean_inc(x_1); -x_41 = l_Lean_Server_DocumentMeta_mkInputContext(x_1); -lean_inc(x_39); -x_42 = lean_apply_2(x_39, x_41, x_40); -x_43 = lean_ctor_get(x_42, 0); +x_34 = l_Lean_Server_DocumentMeta_mkInputContext(x_1); +lean_inc(x_32); +x_35 = lean_apply_2(x_32, x_34, x_33); +x_36 = lean_ctor_get(x_35, 0); +lean_inc(x_36); +x_37 = lean_ctor_get(x_35, 1); +lean_inc(x_37); +lean_dec(x_35); +x_38 = lean_io_promise_result(x_27); +lean_inc(x_24); +lean_inc(x_36); +x_39 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_initializeWorker___lambda__1), 4, 2); +lean_closure_set(x_39, 0, x_36); +lean_closure_set(x_39, 1, x_24); +x_40 = l_Task_Priority_default; +x_41 = 0; +lean_inc(x_38); +x_42 = lean_io_map_task(x_39, x_38, x_40, x_41, x_37); +x_43 = lean_ctor_get(x_42, 1); lean_inc(x_43); -x_44 = lean_ctor_get(x_42, 1); -lean_inc(x_44); lean_dec(x_42); -x_45 = lean_io_promise_result(x_32); -lean_inc(x_29); -lean_inc(x_43); -x_46 = lean_alloc_closure((void*)(l_Lean_Server_FileWorker_initializeWorker___lambda__1), 4, 2); -lean_closure_set(x_46, 0, x_43); -lean_closure_set(x_46, 1, x_29); -x_47 = l_Task_Priority_default; -x_48 = 0; -lean_inc(x_45); -x_49 = lean_io_map_task(x_46, x_45, x_47, x_48, x_44); -x_50 = lean_ctor_get(x_49, 1); +x_44 = lean_alloc_ctor(0, 9, 1); +lean_ctor_set(x_44, 0, x_24); +lean_ctor_set(x_44, 1, x_11); +lean_ctor_set(x_44, 2, x_14); +lean_ctor_set(x_44, 3, x_17); +lean_ctor_set(x_44, 4, x_21); +lean_ctor_set(x_44, 5, x_3); +lean_ctor_set(x_44, 6, x_4); +lean_ctor_set(x_44, 7, x_32); +lean_ctor_set(x_44, 8, x_5); +lean_ctor_set_uint8(x_44, sizeof(void*)*9, x_10); +x_45 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_ReportSnapshotsState_allInfoTrees___default___closed__1; +x_46 = lean_st_mk_ref(x_45, x_43); +x_47 = lean_ctor_get(x_46, 0); +lean_inc(x_47); +x_48 = lean_ctor_get(x_46, 1); +lean_inc(x_48); +lean_dec(x_46); +lean_inc(x_36); +x_49 = l___private_Lean_Server_FileWorker_Utils_0__Lean_Server_FileWorker_mkCmdSnaps(x_36); +x_50 = lean_alloc_ctor(0, 4, 0); +lean_ctor_set(x_50, 0, x_1); +lean_ctor_set(x_50, 1, x_36); +lean_ctor_set(x_50, 2, x_49); +lean_ctor_set(x_50, 3, x_47); +x_51 = l_Lean_Server_FileWorker_CancelToken_new(x_48); +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); +lean_inc(x_52); lean_inc(x_50); -lean_dec(x_49); -x_51 = lean_alloc_ctor(0, 9, 1); -lean_ctor_set(x_51, 0, x_29); -lean_ctor_set(x_51, 1, x_16); -lean_ctor_set(x_51, 2, x_19); -lean_ctor_set(x_51, 3, x_22); -lean_ctor_set(x_51, 4, x_26); -lean_ctor_set(x_51, 5, x_3); -lean_ctor_set(x_51, 6, x_4); -lean_ctor_set(x_51, 7, x_39); -lean_ctor_set(x_51, 8, x_5); -lean_ctor_set_uint8(x_51, sizeof(void*)*9, x_10); -x_52 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_ReportSnapshotsState_allInfoTrees___default___closed__1; -x_53 = lean_st_mk_ref(x_52, x_50); -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); -lean_inc(x_43); -x_56 = l___private_Lean_Server_FileWorker_Utils_0__Lean_Server_FileWorker_mkCmdSnaps(x_43); -x_57 = lean_alloc_ctor(0, 4, 0); -lean_ctor_set(x_57, 0, x_1); -lean_ctor_set(x_57, 1, x_43); -lean_ctor_set(x_57, 2, x_56); -lean_ctor_set(x_57, 3, x_54); -x_58 = l_Lean_Server_FileWorker_CancelToken_new(x_55); -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_59); -lean_inc(x_57); -lean_inc(x_51); -x_61 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots(x_51, x_57, x_59, x_60); -x_62 = !lean_is_exclusive(x_61); -if (x_62 == 0) +lean_inc(x_44); +x_54 = l___private_Lean_Server_FileWorker_0__Lean_Server_FileWorker_reportSnapshots(x_44, x_50, x_52, x_53); +x_55 = !lean_is_exclusive(x_54); +if (x_55 == 0) { -lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; -x_63 = lean_ctor_get(x_61, 0); -x_64 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_64, 0, x_57); -lean_ctor_set(x_64, 1, x_63); -x_65 = lean_box(0); -x_66 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_66, 0, x_64); -lean_ctor_set(x_66, 1, x_59); -lean_ctor_set(x_66, 2, x_45); -lean_ctor_set(x_66, 3, x_65); -lean_ctor_set(x_66, 4, x_24); -lean_ctor_set(x_66, 5, x_24); -x_67 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_67, 0, x_51); -lean_ctor_set(x_67, 1, x_66); -lean_ctor_set(x_61, 0, x_67); -return x_61; +lean_object* x_56; lean_object* x_57; lean_object* x_58; lean_object* x_59; lean_object* x_60; +x_56 = lean_ctor_get(x_54, 0); +x_57 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_57, 0, x_50); +lean_ctor_set(x_57, 1, x_56); +x_58 = lean_box(0); +x_59 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_59, 0, x_57); +lean_ctor_set(x_59, 1, x_52); +lean_ctor_set(x_59, 2, x_38); +lean_ctor_set(x_59, 3, x_58); +lean_ctor_set(x_59, 4, x_19); +lean_ctor_set(x_59, 5, x_19); +x_60 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_60, 0, x_44); +lean_ctor_set(x_60, 1, x_59); +lean_ctor_set(x_54, 0, x_60); +return x_54; } 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; -x_68 = lean_ctor_get(x_61, 0); -x_69 = lean_ctor_get(x_61, 1); -lean_inc(x_69); -lean_inc(x_68); -lean_dec(x_61); -x_70 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_70, 0, x_57); -lean_ctor_set(x_70, 1, x_68); -x_71 = lean_box(0); -x_72 = lean_alloc_ctor(0, 6, 0); -lean_ctor_set(x_72, 0, x_70); -lean_ctor_set(x_72, 1, x_59); -lean_ctor_set(x_72, 2, x_45); -lean_ctor_set(x_72, 3, x_71); -lean_ctor_set(x_72, 4, x_24); -lean_ctor_set(x_72, 5, x_24); -x_73 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_73, 0, x_51); -lean_ctor_set(x_73, 1, x_72); -x_74 = lean_alloc_ctor(0, 2, 0); -lean_ctor_set(x_74, 0, x_73); -lean_ctor_set(x_74, 1, x_69); -return x_74; -} +lean_object* x_61; lean_object* x_62; lean_object* x_63; lean_object* x_64; lean_object* x_65; lean_object* x_66; lean_object* x_67; +x_61 = lean_ctor_get(x_54, 0); +x_62 = lean_ctor_get(x_54, 1); +lean_inc(x_62); +lean_inc(x_61); +lean_dec(x_54); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_50); +lean_ctor_set(x_63, 1, x_61); +x_64 = lean_box(0); +x_65 = lean_alloc_ctor(0, 6, 0); +lean_ctor_set(x_65, 0, x_63); +lean_ctor_set(x_65, 1, x_52); +lean_ctor_set(x_65, 2, x_38); +lean_ctor_set(x_65, 3, x_64); +lean_ctor_set(x_65, 4, x_19); +lean_ctor_set(x_65, 5, x_19); +x_66 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_66, 0, x_44); +lean_ctor_set(x_66, 1, x_65); +x_67 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_67, 0, x_66); +lean_ctor_set(x_67, 1, x_62); +return x_67; } } } @@ -23730,12 +23765,14 @@ l_Lean_Server_FileWorker_setupImports___lambda__1___closed__5 = _init_l_Lean_Ser lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__1___closed__5); l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1 = _init_l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__3___closed__1); -l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1 = _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1(); -lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__4___closed__1); -l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2 = _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__4___closed__2); -l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3 = _init_l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3(); -lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__4___closed__3); +l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1 = _init_l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__5___closed__1); +l_Lean_Server_FileWorker_setupImports___lambda__6___closed__1 = _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__1(); +lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__6___closed__1); +l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2 = _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2(); +lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__6___closed__2); +l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3 = _init_l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3(); +lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___lambda__6___closed__3); l_Lean_Server_FileWorker_setupImports___closed__1 = _init_l_Lean_Server_FileWorker_setupImports___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_setupImports___closed__1); l_Lean_Server_FileWorker_setupImports___closed__2 = _init_l_Lean_Server_FileWorker_setupImports___closed__2(); @@ -23746,8 +23783,6 @@ l_Lean_Server_FileWorker_initializeWorker___lambda__1___closed__1 = _init_l_Lean lean_mark_persistent(l_Lean_Server_FileWorker_initializeWorker___lambda__1___closed__1); l_Lean_Server_FileWorker_initializeWorker___closed__1 = _init_l_Lean_Server_FileWorker_initializeWorker___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_initializeWorker___closed__1); -l_Lean_Server_FileWorker_initializeWorker___closed__2 = _init_l_Lean_Server_FileWorker_initializeWorker___closed__2(); -lean_mark_persistent(l_Lean_Server_FileWorker_initializeWorker___closed__2); l_Lean_Server_FileWorker_sendServerRequest___rarg___closed__1 = _init_l_Lean_Server_FileWorker_sendServerRequest___rarg___closed__1(); lean_mark_persistent(l_Lean_Server_FileWorker_sendServerRequest___rarg___closed__1); l_Lean_Server_FileWorker_handleDidChange___closed__1 = _init_l_Lean_Server_FileWorker_handleDidChange___closed__1(); diff --git a/stage0/stdlib/Lean/Server/Rpc/Deriving.c b/stage0/stdlib/Lean/Server/Rpc/Deriving.c index 352d5564ea25..c99ef5991458 100644 --- a/stage0/stdlib/Lean/Server/Rpc/Deriving.c +++ b/stage0/stdlib/Lean/Server/Rpc/Deriving.c @@ -43,6 +43,7 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__4(size_t, size_t, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___closed__2; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__17; +lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___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_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__3___closed__4; static lean_object* l_Array_forInUnsafe_loop___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___spec__1___closed__13; static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveStructureInstance___lambda__1___closed__104; @@ -170,7 +171,6 @@ static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncod lean_object* l_Lean_Meta_isType(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___lambda__1___closed__12; lean_object* l_Lean_Syntax_node6(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_Server_RpcEncodable_initFn____x40_Lean_Server_Rpc_Deriving___hyg_4899____closed__13; 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*); LEAN_EXPORT lean_object* l_Array_mapMUnsafe_map___at___private_Lean_Server_Rpc_Deriving_0__Lean_Server_RpcEncodable_deriveInductiveInstance___spec__10___boxed(lean_object*, lean_object*, lean_object*, lean_object*); @@ -4635,7 +4635,7 @@ LEAN_EXPORT lean_object* l_Lean_getConstInfoCtor___at___private_Lean_Server_Rpc_ lean_object* x_9; lean_inc(x_2); lean_inc(x_1); -x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__1(x_1, x_2, x_3, x_4, x_5, x_6, x_7, x_8); +x_9 = l_Lean_getConstInfo___at_Lean_Elab_Term_mkConst___spec__3(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; diff --git a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c index dff44282d1a9..271fb9b33e5f 100644 --- a/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c +++ b/stage0/stdlib/Lean/Server/Rpc/RequestHandling.c @@ -120,7 +120,6 @@ static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__15; LEAN_EXPORT lean_object* l_Lean_Server_initFn____x40_Lean_Server_Rpc_RequestHandling___hyg_32_(lean_object*); lean_object* l_Lean_SourceInfo_fromRef(lean_object*, uint8_t); LEAN_EXPORT lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__2___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -lean_object* l_Lean_setEnv___at_Lean_addDecl___spec__3(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*); LEAN_EXPORT uint8_t l_Lean_PersistentHashMap_containsAtAux___at_Lean_Server_registerBuiltinRpcProcedure___spec__7(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_throwError___at_Lean_realizeGlobalConstCore___spec__3(lean_object*, lean_object*, lean_object*, lean_object*); @@ -291,6 +290,7 @@ static lean_object* l_Lean_Server_handleRpcCall___lambda__3___closed__1; lean_object* l_Lean_Meta_InfoCacheKey_instHashable___boxed(lean_object*); static lean_object* l___private_Lean_Server_Rpc_RequestHandling_0__Lean_Server_evalRpcProcedureUnsafe___closed__4; lean_object* lean_array_get_size(lean_object*); +lean_object* l_Lean_setEnv___at_Lean_compileDecl___spec__14(lean_object*, lean_object*, lean_object*, lean_object*); uint8_t lean_nat_dec_le(lean_object*, lean_object*); static lean_object* l_Lean_Server_registerRpcProcedure___lambda__3___closed__31; static lean_object* l_Lean_Server_wrapRpcProcedure___elambda__1___lambda__2___closed__3; @@ -4074,7 +4074,7 @@ lean_inc(x_36); lean_dec(x_34); x_37 = l_Lean_Server_handleRpcCall___lambda__2___closed__1; x_38 = l_Lean_MapDeclarationExtension_insert___rarg(x_37, x_36, x_1, x_7); -x_39 = l_Lean_setEnv___at_Lean_addDecl___spec__3(x_38, x_3, x_4, x_35); +x_39 = l_Lean_setEnv___at_Lean_compileDecl___spec__14(x_38, x_3, x_4, x_35); lean_dec(x_4); lean_dec(x_3); return x_39; diff --git a/stage0/stdlib/Lean/Util/Heartbeats.c b/stage0/stdlib/Lean/Util/Heartbeats.c index 99b04d464d57..adfea5b121fb 100644 --- a/stage0/stdlib/Lean/Util/Heartbeats.c +++ b/stage0/stdlib/Lean/Util/Heartbeats.c @@ -14,33 +14,48 @@ extern "C" { #endif lean_object* l_IO_getNumHeartbeats___boxed(lean_object*); +lean_object* l_Lean_FileMap_toPosition(lean_object*, lean_object*); lean_object* l_Lean_Name_toString(lean_object*, uint8_t); +lean_object* l_Lean_PersistentArray_push___rarg(lean_object*, lean_object*); +lean_object* l_Lean_replaceRef(lean_object*, lean_object*); +lean_object* l_Lean_Syntax_getPos_x3f(lean_object*, uint8_t); +lean_object* l_Lean_Syntax_getTailPos_x3f(lean_object*, uint8_t); +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_MessageData_hasSyntheticSorry(lean_object*); lean_object* lean_io_get_num_heartbeats(lean_object*); LEAN_EXPORT lean_object* l_Lean_getRemainingHeartbeats___boxed(lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_withHeartbeats___rarg___closed__1; +lean_object* lean_st_ref_take(lean_object*, lean_object*); lean_object* lean_nat_div(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_reportOutOfHeartbeats___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg___lambda__3(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); static lean_object* l_Lean_reportOutOfHeartbeats___closed__2; LEAN_EXPORT lean_object* l_Lean_heartbeatsPercent(lean_object*, lean_object*, lean_object*); +lean_object* l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getRemainingHeartbeats(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_reportOutOfHeartbeats(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getInitHeartbeats(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats(lean_object*, lean_object*); static lean_object* l_Lean_reportOutOfHeartbeats___closed__1; +extern lean_object* l_Lean_warningAsError; +uint8_t l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg___lambda__1(lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getMaxHeartbeats___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getInitHeartbeats___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg___lambda__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2; LEAN_EXPORT lean_object* l_Lean_heartbeatsPercent___boxed(lean_object*, lean_object*, lean_object*); lean_object* lean_nat_sub(lean_object*, lean_object*); lean_object* lean_nat_mul(lean_object*, lean_object*); -lean_object* l_Lean_logAt___at_Lean_addDecl___spec__6(lean_object*, lean_object*, uint8_t, lean_object*, lean_object*, lean_object*); +lean_object* lean_st_ref_set(lean_object*, lean_object*, lean_object*); lean_object* lean_string_append(lean_object*, lean_object*); +uint8_t l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(uint8_t, uint8_t); uint8_t lean_nat_dec_le(lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_getMaxHeartbeats(lean_object*, lean_object*, lean_object*); +static lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___boxed(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_withHeartbeats___rarg___lambda__1(lean_object* x_1, lean_object* x_2, lean_object* x_3, lean_object* x_4) { _start: { @@ -304,6 +319,912 @@ lean_dec(x_1); return x_4; } } +static lean_object* _init_l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1() { +_start: +{ +lean_object* x_1; +x_1 = lean_mk_string_from_bytes("", 0); +return x_1; +} +} +static lean_object* _init_l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2() { +_start: +{ +lean_object* x_1; +x_1 = l_Lean_warningAsError; +return x_1; +} +} +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__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; uint8_t x_265; uint8_t x_266; +x_265 = 2; +x_266 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_265); +if (x_266 == 0) +{ +lean_object* x_267; +x_267 = lean_box(0); +x_7 = x_267; +goto block_264; +} +else +{ +lean_object* x_268; uint8_t x_269; +lean_inc(x_2); +x_268 = l_Lean_MessageData_hasSyntheticSorry(x_2); +x_269 = lean_unbox(x_268); +lean_dec(x_268); +if (x_269 == 0) +{ +lean_object* x_270; +x_270 = lean_box(0); +x_7 = x_270; +goto block_264; +} +else +{ +lean_object* x_271; lean_object* x_272; +lean_dec(x_2); +x_271 = lean_box(0); +x_272 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_272, 0, x_271); +lean_ctor_set(x_272, 1, x_6); +return x_272; +} +} +block_264: +{ +uint8_t x_8; lean_object* x_258; uint8_t x_259; uint8_t x_260; +lean_dec(x_7); +x_258 = lean_ctor_get(x_4, 2); +x_259 = 1; +x_260 = l___private_Lean_Message_0__Lean_beqMessageSeverity____x40_Lean_Message___hyg_103_(x_3, x_259); +if (x_260 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +lean_object* x_261; uint8_t x_262; +x_261 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2; +x_262 = l_Lean_Option_get___at___private_Lean_Util_Profile_0__Lean_get__profiler___spec__1(x_258, x_261); +if (x_262 == 0) +{ +x_8 = x_3; +goto block_257; +} +else +{ +uint8_t x_263; +x_263 = 2; +x_8 = x_263; +goto block_257; +} +} +block_257: +{ +lean_object* x_9; lean_object* x_10; lean_object* x_11; lean_object* x_12; lean_object* x_13; lean_object* x_14; uint8_t x_15; lean_object* x_16; lean_object* x_17; +x_9 = lean_ctor_get(x_4, 0); +x_10 = lean_ctor_get(x_4, 1); +x_11 = lean_ctor_get(x_4, 5); +x_12 = lean_ctor_get(x_4, 6); +x_13 = lean_ctor_get(x_4, 7); +x_14 = l_Lean_replaceRef(x_1, x_11); +x_15 = 0; +x_16 = l_Lean_Syntax_getPos_x3f(x_14, x_15); +x_17 = l_Lean_Syntax_getTailPos_x3f(x_14, x_15); +if (lean_obj_tag(x_16) == 0) +{ +if (lean_obj_tag(x_17) == 0) +{ +lean_object* x_18; lean_object* x_19; lean_object* x_20; lean_object* x_21; lean_object* x_22; lean_object* x_23; lean_object* x_24; 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_18 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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_unsigned_to_nat(0u); +x_22 = l_Lean_FileMap_toPosition(x_10, x_21); +lean_inc(x_22); +x_23 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_23, 0, x_22); +lean_inc(x_13); +lean_inc(x_12); +x_24 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_24, 0, x_12); +lean_ctor_set(x_24, 1, x_13); +x_25 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_25, 0, x_24); +lean_ctor_set(x_25, 1, x_19); +x_26 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_27 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_27, 0, x_9); +lean_ctor_set(x_27, 1, x_22); +lean_ctor_set(x_27, 2, x_23); +lean_ctor_set(x_27, 3, x_26); +lean_ctor_set(x_27, 4, x_25); +lean_ctor_set_uint8(x_27, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_27, sizeof(void*)*5 + 1, x_8); +x_28 = lean_st_ref_take(x_5, x_20); +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; uint8_t x_35; +x_32 = lean_ctor_get(x_29, 5); +x_33 = l_Lean_PersistentArray_push___rarg(x_32, x_27); +lean_ctor_set(x_29, 5, x_33); +x_34 = lean_st_ref_set(x_5, x_29, x_30); +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_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_29, 0); +x_42 = lean_ctor_get(x_29, 1); +x_43 = lean_ctor_get(x_29, 2); +x_44 = lean_ctor_get(x_29, 3); +x_45 = lean_ctor_get(x_29, 4); +x_46 = lean_ctor_get(x_29, 5); +x_47 = lean_ctor_get(x_29, 6); +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_inc(x_41); +lean_dec(x_29); +x_48 = l_Lean_PersistentArray_push___rarg(x_46, x_27); +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_44); +lean_ctor_set(x_49, 4, x_45); +lean_ctor_set(x_49, 5, x_48); +lean_ctor_set(x_49, 6, x_47); +x_50 = lean_st_ref_set(x_5, x_49, x_30); +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 +{ +uint8_t x_55; +x_55 = !lean_is_exclusive(x_17); +if (x_55 == 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; lean_object* x_69; uint8_t x_70; +x_56 = lean_ctor_get(x_17, 0); +x_57 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_58 = lean_ctor_get(x_57, 0); +lean_inc(x_58); +x_59 = lean_ctor_get(x_57, 1); +lean_inc(x_59); +lean_dec(x_57); +x_60 = lean_unsigned_to_nat(0u); +x_61 = l_Lean_FileMap_toPosition(x_10, x_60); +x_62 = l_Lean_FileMap_toPosition(x_10, x_56); +lean_dec(x_56); +lean_ctor_set(x_17, 0, x_62); +lean_inc(x_13); +lean_inc(x_12); +x_63 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_63, 0, x_12); +lean_ctor_set(x_63, 1, x_13); +x_64 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_64, 0, x_63); +lean_ctor_set(x_64, 1, x_58); +x_65 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_66 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_66, 0, x_9); +lean_ctor_set(x_66, 1, x_61); +lean_ctor_set(x_66, 2, x_17); +lean_ctor_set(x_66, 3, x_65); +lean_ctor_set(x_66, 4, x_64); +lean_ctor_set_uint8(x_66, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_66, sizeof(void*)*5 + 1, x_8); +x_67 = lean_st_ref_take(x_5, x_59); +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_is_exclusive(x_68); +if (x_70 == 0) +{ +lean_object* x_71; lean_object* x_72; lean_object* x_73; uint8_t x_74; +x_71 = lean_ctor_get(x_68, 5); +x_72 = l_Lean_PersistentArray_push___rarg(x_71, x_66); +lean_ctor_set(x_68, 5, x_72); +x_73 = lean_st_ref_set(x_5, x_68, x_69); +x_74 = !lean_is_exclusive(x_73); +if (x_74 == 0) +{ +lean_object* x_75; lean_object* x_76; +x_75 = lean_ctor_get(x_73, 0); +lean_dec(x_75); +x_76 = lean_box(0); +lean_ctor_set(x_73, 0, x_76); +return x_73; +} +else +{ +lean_object* x_77; lean_object* x_78; lean_object* x_79; +x_77 = lean_ctor_get(x_73, 1); +lean_inc(x_77); +lean_dec(x_73); +x_78 = lean_box(0); +x_79 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_79, 0, x_78); +lean_ctor_set(x_79, 1, x_77); +return x_79; +} +} +else +{ +lean_object* x_80; lean_object* x_81; lean_object* x_82; lean_object* x_83; lean_object* x_84; lean_object* x_85; lean_object* x_86; lean_object* x_87; lean_object* x_88; lean_object* x_89; lean_object* x_90; lean_object* x_91; lean_object* x_92; lean_object* x_93; +x_80 = lean_ctor_get(x_68, 0); +x_81 = lean_ctor_get(x_68, 1); +x_82 = lean_ctor_get(x_68, 2); +x_83 = lean_ctor_get(x_68, 3); +x_84 = lean_ctor_get(x_68, 4); +x_85 = lean_ctor_get(x_68, 5); +x_86 = lean_ctor_get(x_68, 6); +lean_inc(x_86); +lean_inc(x_85); +lean_inc(x_84); +lean_inc(x_83); +lean_inc(x_82); +lean_inc(x_81); +lean_inc(x_80); +lean_dec(x_68); +x_87 = l_Lean_PersistentArray_push___rarg(x_85, x_66); +x_88 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_88, 0, x_80); +lean_ctor_set(x_88, 1, x_81); +lean_ctor_set(x_88, 2, x_82); +lean_ctor_set(x_88, 3, x_83); +lean_ctor_set(x_88, 4, x_84); +lean_ctor_set(x_88, 5, x_87); +lean_ctor_set(x_88, 6, x_86); +x_89 = lean_st_ref_set(x_5, x_88, x_69); +x_90 = lean_ctor_get(x_89, 1); +lean_inc(x_90); +if (lean_is_exclusive(x_89)) { + lean_ctor_release(x_89, 0); + lean_ctor_release(x_89, 1); + x_91 = x_89; +} else { + lean_dec_ref(x_89); + x_91 = lean_box(0); +} +x_92 = lean_box(0); +if (lean_is_scalar(x_91)) { + x_93 = lean_alloc_ctor(0, 2, 0); +} else { + x_93 = x_91; +} +lean_ctor_set(x_93, 0, x_92); +lean_ctor_set(x_93, 1, x_90); +return x_93; +} +} +else +{ +lean_object* x_94; lean_object* x_95; lean_object* x_96; lean_object* x_97; lean_object* x_98; lean_object* x_99; lean_object* x_100; lean_object* x_101; lean_object* x_102; lean_object* x_103; lean_object* x_104; lean_object* x_105; lean_object* x_106; lean_object* x_107; lean_object* x_108; lean_object* x_109; lean_object* x_110; lean_object* x_111; lean_object* x_112; lean_object* x_113; lean_object* x_114; lean_object* x_115; lean_object* x_116; lean_object* x_117; lean_object* x_118; lean_object* x_119; lean_object* x_120; lean_object* x_121; lean_object* x_122; lean_object* x_123; +x_94 = lean_ctor_get(x_17, 0); +lean_inc(x_94); +lean_dec(x_17); +x_95 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_96 = lean_ctor_get(x_95, 0); +lean_inc(x_96); +x_97 = lean_ctor_get(x_95, 1); +lean_inc(x_97); +lean_dec(x_95); +x_98 = lean_unsigned_to_nat(0u); +x_99 = l_Lean_FileMap_toPosition(x_10, x_98); +x_100 = l_Lean_FileMap_toPosition(x_10, x_94); +lean_dec(x_94); +x_101 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_101, 0, x_100); +lean_inc(x_13); +lean_inc(x_12); +x_102 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_102, 0, x_12); +lean_ctor_set(x_102, 1, x_13); +x_103 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_103, 0, x_102); +lean_ctor_set(x_103, 1, x_96); +x_104 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_105 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_105, 0, x_9); +lean_ctor_set(x_105, 1, x_99); +lean_ctor_set(x_105, 2, x_101); +lean_ctor_set(x_105, 3, x_104); +lean_ctor_set(x_105, 4, x_103); +lean_ctor_set_uint8(x_105, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_105, sizeof(void*)*5 + 1, x_8); +x_106 = lean_st_ref_take(x_5, x_97); +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_ctor_get(x_107, 0); +lean_inc(x_109); +x_110 = lean_ctor_get(x_107, 1); +lean_inc(x_110); +x_111 = lean_ctor_get(x_107, 2); +lean_inc(x_111); +x_112 = lean_ctor_get(x_107, 3); +lean_inc(x_112); +x_113 = lean_ctor_get(x_107, 4); +lean_inc(x_113); +x_114 = lean_ctor_get(x_107, 5); +lean_inc(x_114); +x_115 = lean_ctor_get(x_107, 6); +lean_inc(x_115); +if (lean_is_exclusive(x_107)) { + lean_ctor_release(x_107, 0); + lean_ctor_release(x_107, 1); + lean_ctor_release(x_107, 2); + lean_ctor_release(x_107, 3); + lean_ctor_release(x_107, 4); + lean_ctor_release(x_107, 5); + lean_ctor_release(x_107, 6); + x_116 = x_107; +} else { + lean_dec_ref(x_107); + x_116 = lean_box(0); +} +x_117 = l_Lean_PersistentArray_push___rarg(x_114, x_105); +if (lean_is_scalar(x_116)) { + x_118 = lean_alloc_ctor(0, 7, 0); +} else { + x_118 = x_116; +} +lean_ctor_set(x_118, 0, x_109); +lean_ctor_set(x_118, 1, x_110); +lean_ctor_set(x_118, 2, x_111); +lean_ctor_set(x_118, 3, x_112); +lean_ctor_set(x_118, 4, x_113); +lean_ctor_set(x_118, 5, x_117); +lean_ctor_set(x_118, 6, x_115); +x_119 = lean_st_ref_set(x_5, x_118, x_108); +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 +{ +if (lean_obj_tag(x_17) == 0) +{ +uint8_t x_124; +x_124 = !lean_is_exclusive(x_16); +if (x_124 == 0) +{ +lean_object* x_125; lean_object* x_126; lean_object* x_127; lean_object* x_128; lean_object* x_129; lean_object* x_130; lean_object* x_131; lean_object* x_132; lean_object* x_133; lean_object* x_134; lean_object* x_135; lean_object* x_136; uint8_t x_137; +x_125 = lean_ctor_get(x_16, 0); +x_126 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_127 = lean_ctor_get(x_126, 0); +lean_inc(x_127); +x_128 = lean_ctor_get(x_126, 1); +lean_inc(x_128); +lean_dec(x_126); +x_129 = l_Lean_FileMap_toPosition(x_10, x_125); +lean_dec(x_125); +lean_inc(x_129); +lean_ctor_set(x_16, 0, x_129); +lean_inc(x_13); +lean_inc(x_12); +x_130 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_130, 0, x_12); +lean_ctor_set(x_130, 1, x_13); +x_131 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_131, 0, x_130); +lean_ctor_set(x_131, 1, x_127); +x_132 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_133 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_133, 0, x_9); +lean_ctor_set(x_133, 1, x_129); +lean_ctor_set(x_133, 2, x_16); +lean_ctor_set(x_133, 3, x_132); +lean_ctor_set(x_133, 4, x_131); +lean_ctor_set_uint8(x_133, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_133, sizeof(void*)*5 + 1, x_8); +x_134 = lean_st_ref_take(x_5, x_128); +x_135 = lean_ctor_get(x_134, 0); +lean_inc(x_135); +x_136 = lean_ctor_get(x_134, 1); +lean_inc(x_136); +lean_dec(x_134); +x_137 = !lean_is_exclusive(x_135); +if (x_137 == 0) +{ +lean_object* x_138; lean_object* x_139; lean_object* x_140; uint8_t x_141; +x_138 = lean_ctor_get(x_135, 5); +x_139 = l_Lean_PersistentArray_push___rarg(x_138, x_133); +lean_ctor_set(x_135, 5, x_139); +x_140 = lean_st_ref_set(x_5, x_135, x_136); +x_141 = !lean_is_exclusive(x_140); +if (x_141 == 0) +{ +lean_object* x_142; lean_object* x_143; +x_142 = lean_ctor_get(x_140, 0); +lean_dec(x_142); +x_143 = lean_box(0); +lean_ctor_set(x_140, 0, x_143); +return x_140; +} +else +{ +lean_object* x_144; lean_object* x_145; lean_object* x_146; +x_144 = lean_ctor_get(x_140, 1); +lean_inc(x_144); +lean_dec(x_140); +x_145 = lean_box(0); +x_146 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_146, 0, x_145); +lean_ctor_set(x_146, 1, x_144); +return x_146; +} +} +else +{ +lean_object* x_147; lean_object* x_148; lean_object* x_149; lean_object* x_150; lean_object* x_151; lean_object* x_152; lean_object* x_153; lean_object* x_154; lean_object* x_155; lean_object* x_156; lean_object* x_157; lean_object* x_158; lean_object* x_159; lean_object* x_160; +x_147 = lean_ctor_get(x_135, 0); +x_148 = lean_ctor_get(x_135, 1); +x_149 = lean_ctor_get(x_135, 2); +x_150 = lean_ctor_get(x_135, 3); +x_151 = lean_ctor_get(x_135, 4); +x_152 = lean_ctor_get(x_135, 5); +x_153 = lean_ctor_get(x_135, 6); +lean_inc(x_153); +lean_inc(x_152); +lean_inc(x_151); +lean_inc(x_150); +lean_inc(x_149); +lean_inc(x_148); +lean_inc(x_147); +lean_dec(x_135); +x_154 = l_Lean_PersistentArray_push___rarg(x_152, x_133); +x_155 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_155, 0, x_147); +lean_ctor_set(x_155, 1, x_148); +lean_ctor_set(x_155, 2, x_149); +lean_ctor_set(x_155, 3, x_150); +lean_ctor_set(x_155, 4, x_151); +lean_ctor_set(x_155, 5, x_154); +lean_ctor_set(x_155, 6, x_153); +x_156 = lean_st_ref_set(x_5, x_155, x_136); +x_157 = lean_ctor_get(x_156, 1); +lean_inc(x_157); +if (lean_is_exclusive(x_156)) { + lean_ctor_release(x_156, 0); + lean_ctor_release(x_156, 1); + x_158 = x_156; +} else { + lean_dec_ref(x_156); + x_158 = lean_box(0); +} +x_159 = lean_box(0); +if (lean_is_scalar(x_158)) { + x_160 = lean_alloc_ctor(0, 2, 0); +} else { + x_160 = x_158; +} +lean_ctor_set(x_160, 0, x_159); +lean_ctor_set(x_160, 1, x_157); +return x_160; +} +} +else +{ +lean_object* x_161; lean_object* x_162; lean_object* x_163; lean_object* x_164; lean_object* x_165; lean_object* x_166; lean_object* x_167; lean_object* x_168; lean_object* x_169; lean_object* x_170; lean_object* x_171; lean_object* x_172; lean_object* x_173; lean_object* x_174; lean_object* x_175; lean_object* x_176; lean_object* x_177; lean_object* x_178; lean_object* x_179; lean_object* x_180; lean_object* x_181; lean_object* x_182; lean_object* x_183; lean_object* x_184; lean_object* x_185; lean_object* x_186; lean_object* x_187; lean_object* x_188; +x_161 = lean_ctor_get(x_16, 0); +lean_inc(x_161); +lean_dec(x_16); +x_162 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_163 = lean_ctor_get(x_162, 0); +lean_inc(x_163); +x_164 = lean_ctor_get(x_162, 1); +lean_inc(x_164); +lean_dec(x_162); +x_165 = l_Lean_FileMap_toPosition(x_10, x_161); +lean_dec(x_161); +lean_inc(x_165); +x_166 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_166, 0, x_165); +lean_inc(x_13); +lean_inc(x_12); +x_167 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_167, 0, x_12); +lean_ctor_set(x_167, 1, x_13); +x_168 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_168, 0, x_167); +lean_ctor_set(x_168, 1, x_163); +x_169 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_170 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_170, 0, x_9); +lean_ctor_set(x_170, 1, x_165); +lean_ctor_set(x_170, 2, x_166); +lean_ctor_set(x_170, 3, x_169); +lean_ctor_set(x_170, 4, x_168); +lean_ctor_set_uint8(x_170, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_170, sizeof(void*)*5 + 1, x_8); +x_171 = lean_st_ref_take(x_5, x_164); +x_172 = lean_ctor_get(x_171, 0); +lean_inc(x_172); +x_173 = lean_ctor_get(x_171, 1); +lean_inc(x_173); +lean_dec(x_171); +x_174 = lean_ctor_get(x_172, 0); +lean_inc(x_174); +x_175 = lean_ctor_get(x_172, 1); +lean_inc(x_175); +x_176 = lean_ctor_get(x_172, 2); +lean_inc(x_176); +x_177 = lean_ctor_get(x_172, 3); +lean_inc(x_177); +x_178 = lean_ctor_get(x_172, 4); +lean_inc(x_178); +x_179 = lean_ctor_get(x_172, 5); +lean_inc(x_179); +x_180 = lean_ctor_get(x_172, 6); +lean_inc(x_180); +if (lean_is_exclusive(x_172)) { + lean_ctor_release(x_172, 0); + lean_ctor_release(x_172, 1); + lean_ctor_release(x_172, 2); + lean_ctor_release(x_172, 3); + lean_ctor_release(x_172, 4); + lean_ctor_release(x_172, 5); + lean_ctor_release(x_172, 6); + x_181 = x_172; +} else { + lean_dec_ref(x_172); + x_181 = lean_box(0); +} +x_182 = l_Lean_PersistentArray_push___rarg(x_179, x_170); +if (lean_is_scalar(x_181)) { + x_183 = lean_alloc_ctor(0, 7, 0); +} else { + x_183 = x_181; +} +lean_ctor_set(x_183, 0, x_174); +lean_ctor_set(x_183, 1, x_175); +lean_ctor_set(x_183, 2, x_176); +lean_ctor_set(x_183, 3, x_177); +lean_ctor_set(x_183, 4, x_178); +lean_ctor_set(x_183, 5, x_182); +lean_ctor_set(x_183, 6, x_180); +x_184 = lean_st_ref_set(x_5, x_183, x_173); +x_185 = lean_ctor_get(x_184, 1); +lean_inc(x_185); +if (lean_is_exclusive(x_184)) { + lean_ctor_release(x_184, 0); + lean_ctor_release(x_184, 1); + x_186 = x_184; +} else { + lean_dec_ref(x_184); + x_186 = lean_box(0); +} +x_187 = lean_box(0); +if (lean_is_scalar(x_186)) { + x_188 = lean_alloc_ctor(0, 2, 0); +} else { + x_188 = x_186; +} +lean_ctor_set(x_188, 0, x_187); +lean_ctor_set(x_188, 1, x_185); +return x_188; +} +} +else +{ +lean_object* x_189; uint8_t x_190; +x_189 = lean_ctor_get(x_16, 0); +lean_inc(x_189); +lean_dec(x_16); +x_190 = !lean_is_exclusive(x_17); +if (x_190 == 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; 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; uint8_t x_204; +x_191 = lean_ctor_get(x_17, 0); +x_192 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +x_193 = lean_ctor_get(x_192, 0); +lean_inc(x_193); +x_194 = lean_ctor_get(x_192, 1); +lean_inc(x_194); +lean_dec(x_192); +x_195 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_196 = l_Lean_FileMap_toPosition(x_10, x_191); +lean_dec(x_191); +lean_ctor_set(x_17, 0, x_196); +lean_inc(x_13); +lean_inc(x_12); +x_197 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_197, 0, x_12); +lean_ctor_set(x_197, 1, x_13); +x_198 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_198, 0, x_197); +lean_ctor_set(x_198, 1, x_193); +x_199 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_200 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_200, 0, x_9); +lean_ctor_set(x_200, 1, x_195); +lean_ctor_set(x_200, 2, x_17); +lean_ctor_set(x_200, 3, x_199); +lean_ctor_set(x_200, 4, x_198); +lean_ctor_set_uint8(x_200, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_200, sizeof(void*)*5 + 1, x_8); +x_201 = lean_st_ref_take(x_5, x_194); +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_is_exclusive(x_202); +if (x_204 == 0) +{ +lean_object* x_205; lean_object* x_206; lean_object* x_207; uint8_t x_208; +x_205 = lean_ctor_get(x_202, 5); +x_206 = l_Lean_PersistentArray_push___rarg(x_205, x_200); +lean_ctor_set(x_202, 5, x_206); +x_207 = lean_st_ref_set(x_5, x_202, x_203); +x_208 = !lean_is_exclusive(x_207); +if (x_208 == 0) +{ +lean_object* x_209; lean_object* x_210; +x_209 = lean_ctor_get(x_207, 0); +lean_dec(x_209); +x_210 = lean_box(0); +lean_ctor_set(x_207, 0, x_210); +return x_207; +} +else +{ +lean_object* x_211; lean_object* x_212; lean_object* x_213; +x_211 = lean_ctor_get(x_207, 1); +lean_inc(x_211); +lean_dec(x_207); +x_212 = lean_box(0); +x_213 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_213, 0, x_212); +lean_ctor_set(x_213, 1, x_211); +return x_213; +} +} +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; +x_214 = lean_ctor_get(x_202, 0); +x_215 = lean_ctor_get(x_202, 1); +x_216 = lean_ctor_get(x_202, 2); +x_217 = lean_ctor_get(x_202, 3); +x_218 = lean_ctor_get(x_202, 4); +x_219 = lean_ctor_get(x_202, 5); +x_220 = lean_ctor_get(x_202, 6); +lean_inc(x_220); +lean_inc(x_219); +lean_inc(x_218); +lean_inc(x_217); +lean_inc(x_216); +lean_inc(x_215); +lean_inc(x_214); +lean_dec(x_202); +x_221 = l_Lean_PersistentArray_push___rarg(x_219, x_200); +x_222 = lean_alloc_ctor(0, 7, 0); +lean_ctor_set(x_222, 0, x_214); +lean_ctor_set(x_222, 1, x_215); +lean_ctor_set(x_222, 2, x_216); +lean_ctor_set(x_222, 3, x_217); +lean_ctor_set(x_222, 4, x_218); +lean_ctor_set(x_222, 5, x_221); +lean_ctor_set(x_222, 6, x_220); +x_223 = lean_st_ref_set(x_5, x_222, x_203); +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_223); + x_225 = lean_box(0); +} +x_226 = lean_box(0); +if (lean_is_scalar(x_225)) { + x_227 = lean_alloc_ctor(0, 2, 0); +} else { + x_227 = x_225; +} +lean_ctor_set(x_227, 0, x_226); +lean_ctor_set(x_227, 1, x_224); +return x_227; +} +} +else +{ +lean_object* x_228; lean_object* x_229; lean_object* x_230; lean_object* x_231; lean_object* x_232; lean_object* x_233; lean_object* x_234; lean_object* x_235; lean_object* x_236; lean_object* x_237; lean_object* x_238; lean_object* x_239; lean_object* x_240; lean_object* x_241; lean_object* x_242; lean_object* x_243; lean_object* x_244; lean_object* x_245; lean_object* x_246; lean_object* x_247; lean_object* x_248; lean_object* x_249; lean_object* x_250; lean_object* x_251; lean_object* x_252; lean_object* x_253; lean_object* x_254; lean_object* x_255; lean_object* x_256; +x_228 = lean_ctor_get(x_17, 0); +lean_inc(x_228); +lean_dec(x_17); +x_229 = l_Lean_addMessageContextPartial___at_Lean_Core_instAddMessageContextCoreM___spec__1(x_2, x_4, x_5, x_6); +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 = l_Lean_FileMap_toPosition(x_10, x_189); +lean_dec(x_189); +x_233 = l_Lean_FileMap_toPosition(x_10, x_228); +lean_dec(x_228); +x_234 = lean_alloc_ctor(1, 1, 0); +lean_ctor_set(x_234, 0, x_233); +lean_inc(x_13); +lean_inc(x_12); +x_235 = lean_alloc_ctor(0, 2, 0); +lean_ctor_set(x_235, 0, x_12); +lean_ctor_set(x_235, 1, x_13); +x_236 = lean_alloc_ctor(4, 2, 0); +lean_ctor_set(x_236, 0, x_235); +lean_ctor_set(x_236, 1, x_230); +x_237 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1; +lean_inc(x_9); +x_238 = lean_alloc_ctor(0, 5, 2); +lean_ctor_set(x_238, 0, x_9); +lean_ctor_set(x_238, 1, x_232); +lean_ctor_set(x_238, 2, x_234); +lean_ctor_set(x_238, 3, x_237); +lean_ctor_set(x_238, 4, x_236); +lean_ctor_set_uint8(x_238, sizeof(void*)*5, x_15); +lean_ctor_set_uint8(x_238, sizeof(void*)*5 + 1, x_8); +x_239 = lean_st_ref_take(x_5, x_231); +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 = lean_ctor_get(x_240, 0); +lean_inc(x_242); +x_243 = lean_ctor_get(x_240, 1); +lean_inc(x_243); +x_244 = lean_ctor_get(x_240, 2); +lean_inc(x_244); +x_245 = lean_ctor_get(x_240, 3); +lean_inc(x_245); +x_246 = lean_ctor_get(x_240, 4); +lean_inc(x_246); +x_247 = lean_ctor_get(x_240, 5); +lean_inc(x_247); +x_248 = lean_ctor_get(x_240, 6); +lean_inc(x_248); +if (lean_is_exclusive(x_240)) { + lean_ctor_release(x_240, 0); + lean_ctor_release(x_240, 1); + lean_ctor_release(x_240, 2); + lean_ctor_release(x_240, 3); + lean_ctor_release(x_240, 4); + lean_ctor_release(x_240, 5); + lean_ctor_release(x_240, 6); + x_249 = x_240; +} else { + lean_dec_ref(x_240); + x_249 = lean_box(0); +} +x_250 = l_Lean_PersistentArray_push___rarg(x_247, x_238); +if (lean_is_scalar(x_249)) { + x_251 = lean_alloc_ctor(0, 7, 0); +} else { + x_251 = x_249; +} +lean_ctor_set(x_251, 0, x_242); +lean_ctor_set(x_251, 1, x_243); +lean_ctor_set(x_251, 2, x_244); +lean_ctor_set(x_251, 3, x_245); +lean_ctor_set(x_251, 4, x_246); +lean_ctor_set(x_251, 5, x_250); +lean_ctor_set(x_251, 6, x_248); +x_252 = lean_st_ref_set(x_5, x_251, x_241); +x_253 = lean_ctor_get(x_252, 1); +lean_inc(x_253); +if (lean_is_exclusive(x_252)) { + lean_ctor_release(x_252, 0); + lean_ctor_release(x_252, 1); + x_254 = x_252; +} else { + lean_dec_ref(x_252); + x_254 = lean_box(0); +} +x_255 = lean_box(0); +if (lean_is_scalar(x_254)) { + x_256 = lean_alloc_ctor(0, 2, 0); +} else { + x_256 = x_254; +} +lean_ctor_set(x_256, 0, x_255); +lean_ctor_set(x_256, 1, x_253); +return x_256; +} +} +} +} +} +} +} static lean_object* _init_l_Lean_reportOutOfHeartbeats___closed__1() { _start: { @@ -333,9 +1254,11 @@ x_9 = lean_ctor_get(x_7, 0); x_10 = lean_ctor_get(x_7, 1); x_11 = lean_nat_dec_le(x_3, x_9); lean_dec(x_9); +lean_dec(x_3); if (x_11 == 0) { lean_object* x_12; +lean_dec(x_2); lean_dec(x_1); x_12 = lean_box(0); lean_ctor_set(x_7, 0, x_12); @@ -357,7 +1280,8 @@ lean_ctor_set(x_19, 0, x_18); x_20 = lean_alloc_ctor(0, 1, 0); lean_ctor_set(x_20, 0, x_19); x_21 = 0; -x_22 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_2, x_20, x_21, x_4, x_5, x_10); +x_22 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1(x_2, x_20, x_21, x_4, x_5, x_10); +lean_dec(x_2); return x_22; } } @@ -371,9 +1295,11 @@ lean_inc(x_23); lean_dec(x_7); x_25 = lean_nat_dec_le(x_3, x_23); lean_dec(x_23); +lean_dec(x_3); if (x_25 == 0) { lean_object* x_26; lean_object* x_27; +lean_dec(x_2); lean_dec(x_1); x_26 = lean_box(0); x_27 = lean_alloc_ctor(0, 2, 0); @@ -396,12 +1322,26 @@ 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 = 0; -x_37 = l_Lean_logAt___at_Lean_addDecl___spec__6(x_2, x_35, x_36, x_4, x_5, x_24); +x_37 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1(x_2, x_35, x_36, x_4, x_5, x_24); +lean_dec(x_2); return x_37; } } } } +LEAN_EXPORT lean_object* l_Lean_logAt___at_Lean_reportOutOfHeartbeats___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: +{ +uint8_t x_7; lean_object* x_8; +x_7 = lean_unbox(x_3); +lean_dec(x_3); +x_8 = l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1(x_1, x_2, x_7, x_4, x_5, x_6); +lean_dec(x_5); +lean_dec(x_4); +lean_dec(x_1); +return x_8; +} +} LEAN_EXPORT lean_object* l_Lean_reportOutOfHeartbeats___boxed(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: { @@ -409,8 +1349,6 @@ lean_object* x_7; x_7 = l_Lean_reportOutOfHeartbeats(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; } } @@ -425,6 +1363,10 @@ if (lean_io_result_is_error(res)) return res; lean_dec_ref(res); l_Lean_withHeartbeats___rarg___closed__1 = _init_l_Lean_withHeartbeats___rarg___closed__1(); lean_mark_persistent(l_Lean_withHeartbeats___rarg___closed__1); +l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1 = _init_l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1(); +lean_mark_persistent(l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__1); +l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2 = _init_l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2(); +lean_mark_persistent(l_Lean_logAt___at_Lean_reportOutOfHeartbeats___spec__1___closed__2); l_Lean_reportOutOfHeartbeats___closed__1 = _init_l_Lean_reportOutOfHeartbeats___closed__1(); lean_mark_persistent(l_Lean_reportOutOfHeartbeats___closed__1); l_Lean_reportOutOfHeartbeats___closed__2 = _init_l_Lean_reportOutOfHeartbeats___closed__2(); diff --git a/stage0/stdlib/Lean/Widget/UserWidget.c b/stage0/stdlib/Lean/Widget/UserWidget.c index d06611b52074..44ff8dc51665 100644 --- a/stage0/stdlib/Lean/Widget/UserWidget.c +++ b/stage0/stdlib/Lean/Widget/UserWidget.c @@ -269,6 +269,7 @@ LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__4___boxed(lean_ LEAN_EXPORT lean_object* l_Lean_Elab_liftMacroM___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__3___lambda__2(lean_object*, lean_object*, lean_object*, lean_object*); lean_object* l_Lean_Json_opt___at___private_Lean_Widget_InteractiveCode_0__Lean_Widget_Lean_Widget_SubexprInfo_toJsonRpcEncodablePacket____x40_Lean_Widget_InteractiveCode___hyg_549____spec__1(lean_object*, lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonWidgetSource____x40_Lean_Widget_UserWidget___hyg_1007____closed__7; +uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_fromJsonGetWidgetSourceParams____x40_Lean_Widget_UserWidget___hyg_855_(lean_object*); static lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_Lean_Widget_GetWidgetsResponse_fromJsonRpcEncodablePacket____x40_Lean_Widget_UserWidget___hyg_5282____closed__7; static lean_object* l_Lean_Widget_showPanelWidgetsCmd___closed__7; @@ -625,7 +626,6 @@ LEAN_EXPORT lean_object* l_Lean_Widget_getWidgetSource___lambda__1(lean_object*, static lean_object* l_Lean_Widget_initFn____x40_Lean_Widget_UserWidget___hyg_232____lambda__8___closed__2; LEAN_EXPORT lean_object* l___private_Lean_Widget_UserWidget_0__Lean_Widget_widgetAttrImpl___elambda__1___rarg___boxed(lean_object*, lean_object*, lean_object*); LEAN_EXPORT lean_object* l_Lean_throwError___at_Lean_Widget_elabShowPanelWidgetsCmd___spec__5(lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*, lean_object*); -uint8_t l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(uint8_t, uint8_t); LEAN_EXPORT lean_object* l_Lean_Widget_addPanelWidgetScoped___rarg(lean_object*, lean_object*, lean_object*, uint64_t, lean_object*); lean_object* l_Lean_Syntax_node1(lean_object*, lean_object*, lean_object*); uint8_t l_Lean_Expr_isConstOf(lean_object*, lean_object*); @@ -7087,7 +7087,7 @@ x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); x_15 = 0; -x_16 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_15); +x_16 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(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; lean_object* x_22; lean_object* x_23; uint8_t x_24; @@ -7521,7 +7521,7 @@ x_14 = lean_ctor_get(x_13, 1); lean_inc(x_14); lean_dec(x_13); x_15 = 0; -x_16 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_183_(x_4, x_15); +x_16 = l___private_Lean_Attributes_0__Lean_beqAttributeKind____x40_Lean_Attributes___hyg_153_(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; lean_object* x_22; lean_object* x_23; uint8_t x_24; From 228ff58f3afc614569ce0f922a02ea4b36a4dbde Mon Sep 17 00:00:00 2001 From: JovanGerb <56355248+JovanGerb@users.noreply.github.com> Date: Fri, 10 May 2024 21:35:21 +0100 Subject: [PATCH 29/44] chore: remove duplicate check (#4126) --- src/Lean/Elab/PreDefinition/WF/Main.lean | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/src/Lean/Elab/PreDefinition/WF/Main.lean b/src/Lean/Elab/PreDefinition/WF/Main.lean index 97a7560e8413..636eda87a64e 100644 --- a/src/Lean/Elab/PreDefinition/WF/Main.lean +++ b/src/Lean/Elab/PreDefinition/WF/Main.lean @@ -151,7 +151,7 @@ def wfRecursion (preDefs : Array PreDefinition) : TermElabM Unit := do applyAttributesOf #[preDef] AttributeApplicationTime.afterCompilation -- Unless the user asks for something else, mark the definition as irreducible unless preDef.modifiers.attrs.any fun a => - a.name = `semireducible || a.name = `reducible || a.name = `semireducible do + a.name = `reducible || a.name = `semireducible do setIrreducibleAttribute preDef.declName builtin_initialize registerTraceClass `Elab.definition.wf From 6c6b56e7fc99ebe0fbe67716d5322d853fc1c7d7 Mon Sep 17 00:00:00 2001 From: Arthur Adjedj Date: Fri, 10 May 2024 22:36:08 +0200 Subject: [PATCH 30/44] fix: revert "monadic generalization of FindExpr" (#4125) This reverts commit 706a4cfd73a23ab85c71691587c77b40ddbf149c introduced in #3970 As explained in #4124, `findM?` can become a footgun if used in monads which induce side-effects such as caching. This PR removes that function, and fixes the code introduced by #3398 for which the function was first added. cc @semorrison. --- src/Lean/Meta/Injective.lean | 16 ++++++------- src/Lean/Util/FindExpr.lean | 44 ++++++++++++++++-------------------- 2 files changed, 27 insertions(+), 33 deletions(-) diff --git a/src/Lean/Meta/Injective.lean b/src/Lean/Meta/Injective.lean index 7fefdd21234f..ad098581e410 100644 --- a/src/Lean/Meta/Injective.lean +++ b/src/Lean/Meta/Injective.lean @@ -52,13 +52,13 @@ def elimOptParam (type : Expr) : CoreM Expr := do Instead of checking the type of every subterm, we only need to check the type of free variables, since free variables introduced in the constructor may only appear in the type of other free variables introduced after them. -/ -def occursOrInType (e : Expr) (t : Expr) : MetaM Bool := do - let_fun f (s : Expr) := do - if !s.isFVar then - return s == e - let ty ← inferType s - return s == e || e.occurs ty - return (← t.findM? f).isSome +def occursOrInType (lctx : LocalContext) (e : Expr) (t : Expr) : Bool := + t.find? go |>.isSome +where + go s := Id.run do + let .fvar fvarId := s | s == e + let some decl := lctx.find? fvarId | s == e + return s == e || e.occurs decl.type private partial def mkInjectiveTheoremTypeCore? (ctorVal : ConstructorVal) (useEq : Bool) : MetaM (Option Expr) := do let us := ctorVal.levelParams.map mkLevelParam @@ -87,7 +87,7 @@ private partial def mkInjectiveTheoremTypeCore? (ctorVal : ConstructorVal) (useE match (← whnf type) with | Expr.forallE n d b _ => let arg1 := args1.get ⟨i, h⟩ - if ← occursOrInType arg1 resultType then + if occursOrInType (← getLCtx) arg1 resultType then mkArgs2 (i + 1) (b.instantiate1 arg1) (args2.push arg1) args2New else withLocalDecl n (if useEq then BinderInfo.default else BinderInfo.implicit) d fun arg2 => diff --git a/src/Lean/Util/FindExpr.lean b/src/Lean/Util/FindExpr.lean index e9788ea76bf1..a93765a3cb8c 100644 --- a/src/Lean/Util/FindExpr.lean +++ b/src/Lean/Util/FindExpr.lean @@ -11,17 +11,17 @@ namespace Lean namespace Expr namespace FindImpl -unsafe abbrev FindM (m) := StateT (PtrSet Expr) m +unsafe abbrev FindM := StateT (PtrSet Expr) Id -@[inline] unsafe def checkVisited [Monad m] (e : Expr) : OptionT (FindM m) Unit := do +@[inline] unsafe def checkVisited (e : Expr) : OptionT FindM Unit := do if (← get).contains e then failure modify fun s => s.insert e -unsafe def findM? [Monad m] (p : Expr → m Bool) (e : Expr) : OptionT (FindM m) Expr := +unsafe def findM? (p : Expr → Bool) (e : Expr) : OptionT FindM Expr := let rec visit (e : Expr) := do checkVisited e - if ← p e then + if p e then pure e else match e with | .forallE _ d b _ => visit d <|> visit b @@ -33,35 +33,29 @@ unsafe def findM? [Monad m] (p : Expr → m Bool) (e : Expr) : OptionT (FindM m) | _ => failure visit e -unsafe def findUnsafeM? {m} [Monad m] (p : Expr → m Bool) (e : Expr) : m (Option Expr) := - findM? p e |>.run' mkPtrSet - -@[inline] unsafe def findUnsafe? (p : Expr → Bool) (e : Expr) : Option Expr := findUnsafeM? (m := Id) p e +unsafe def findUnsafe? (p : Expr → Bool) (e : Expr) : Option Expr := + Id.run <| findM? p e |>.run' mkPtrSet end FindImpl -@[implemented_by FindImpl.findUnsafeM?] -/- This is a reference implementation for the unsafe one above -/ -def findM? [Monad m] (p : Expr → m Bool) (e : Expr) : m (Option Expr) := do - if ← p e then - return some e - else match e with - | .forallE _ d b _ => findM? p d <||> findM? p b - | .lam _ d b _ => findM? p d <||> findM? p b - | .mdata _ b => findM? p b - | .letE _ t v b _ => findM? p t <||> findM? p v <||> findM? p b - | .app f a => findM? p f <||> findM? p a - | .proj _ _ b => findM? p b - | _ => pure none - @[implemented_by FindImpl.findUnsafe?] -def find? (p : Expr → Bool) (e : Expr) : Option Expr := findM? (m := Id) p e +def find? (p : Expr → Bool) (e : Expr) : Option Expr := + /- This is a reference implementation for the unsafe one above -/ + if p e then + some e + else match e with + | .forallE _ d b _ => find? p d <|> find? p b + | .lam _ d b _ => find? p d <|> find? p b + | .mdata _ b => find? p b + | .letE _ t v b _ => find? p t <|> find? p v <|> find? p b + | .app f a => find? p f <|> find? p a + | .proj _ _ b => find? p b + | _ => none /-- Return true if `e` occurs in `t` -/ def occurs (e : Expr) (t : Expr) : Bool := (t.find? fun s => s == e).isSome - /-- Return type for `findExt?` function argument. -/ @@ -72,7 +66,7 @@ inductive FindStep where namespace FindExtImpl -unsafe def findM? (p : Expr → FindStep) (e : Expr) : OptionT (FindImpl.FindM Id) Expr := +unsafe def findM? (p : Expr → FindStep) (e : Expr) : OptionT FindImpl.FindM Expr := visit e where visitApp (e : Expr) := From a6d186a81d4a4b5991b568cf016b1f471646b79a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Fri, 10 May 2024 23:55:26 +0200 Subject: [PATCH 31/44] fix: add `checkSystem` and `withIncRecDepth` to `withAutoBoundImplicit` (#4128) Fix stack overflow crash. Closes #4117 The fix can be improved: we could try to avoid creating hundreds of auto implicits before failing. --- src/Lean/Elab/Term.lean | 3 ++- tests/lean/4117.lean | 1 + tests/lean/4117.lean.expected.out | 1 + tests/lean/run/simpDiag.lean | 4 ++-- 4 files changed, 6 insertions(+), 3 deletions(-) create mode 100644 tests/lean/4117.lean create mode 100644 tests/lean/4117.lean.expected.out diff --git a/src/Lean/Elab/Term.lean b/src/Lean/Elab/Term.lean index ee1014f94442..91d631d98ee0 100644 --- a/src/Lean/Elab/Term.lean +++ b/src/Lean/Elab/Term.lean @@ -1523,7 +1523,8 @@ partial def withAutoBoundImplicit (k : TermElabM α) : TermElabM α := do let flag := autoImplicit.get (← getOptions) if flag then withReader (fun ctx => { ctx with autoBoundImplicit := flag, autoBoundImplicits := {} }) do - let rec loop (s : SavedState) : TermElabM α := do + let rec loop (s : SavedState) : TermElabM α := withIncRecDepth do + checkSystem "auto-implicit" try k catch diff --git a/tests/lean/4117.lean b/tests/lean/4117.lean new file mode 100644 index 000000000000..b60ae256cdbc --- /dev/null +++ b/tests/lean/4117.lean @@ -0,0 +1 @@ +structure D:=(t:A)(c N:={s with} 0 diff --git a/tests/lean/4117.lean.expected.out b/tests/lean/4117.lean.expected.out new file mode 100644 index 000000000000..b47ea97b7859 --- /dev/null +++ b/tests/lean/4117.lean.expected.out @@ -0,0 +1 @@ +4117.lean:2:0: error: unexpected end of input; expected ')' diff --git a/tests/lean/run/simpDiag.lean b/tests/lean/run/simpDiag.lean index 4c12eeeaf7c3..91426ff5c785 100644 --- a/tests/lean/run/simpDiag.lean +++ b/tests/lean/run/simpDiag.lean @@ -35,8 +35,8 @@ info: [simp] used theorems (max: 1201, num: 3): ⏎ Nat.reduceAdd (builtin simproc) ↦ 771 ⏎ - ack.eq_1 ↦ 768[simp] tried theorems (max: 1974, num: 2): - ack.eq_3 ↦ 1974, succeeded: 1201 + ack.eq_1 ↦ 768[simp] tried theorems (max: 1973, num: 2): + ack.eq_3 ↦ 1973, succeeded: 1201 ⏎ ack.eq_1 ↦ 768, succeeded: 768use `set_option diagnostics.threshold ` to control threshold for reporting counters --- From e237e12478cd101915a232bfcac258ce014a0f21 Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 May 2024 00:34:09 +0200 Subject: [PATCH 32/44] refactor: add `tryCatchRuntimeEx` combinator (#4129) see #4079 --- src/Lean/CoreM.lean | 42 ++++++++++++++--------------- src/Lean/Meta/SynthInstance.lean | 9 +++---- src/Lean/Meta/Tactic/Simp/Main.lean | 36 ++++++++++++------------- 3 files changed, 42 insertions(+), 45 deletions(-) diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index c92760026cdd..fc1923be0627 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -79,12 +79,6 @@ structure Context where maxHeartbeats : Nat := getMaxHeartbeats options currMacroScope : MacroScope := firstFrontendMacroScope /-- - If `catchRuntimeEx = false`, then given `try x catch ex => h ex`, - an runtime exception occurring in `x` is not handled by `h`. - Recall that runtime exceptions are `maxRecDepth` or `maxHeartbeats`. - -/ - catchRuntimeEx : Bool := false - /-- If `diag := true`, different parts of the system collect diagnostics. Use the `set_option diag true` to set it to true. -/ @@ -415,30 +409,36 @@ in these monads, but on `CommandElabM`. See issues #2775 and #2744 as well as `M try x catch ex => - if ex.isRuntime && !(← read).catchRuntimeEx then - throw ex + if ex.isRuntime then + throw ex -- We should use `tryCatchRuntimeEx` for catching runtime exceptions else h ex +@[inline] protected def Core.tryCatchRuntimeEx (x : CoreM α) (h : Exception → CoreM α) : CoreM α := do + try + x + catch ex => + h ex + instance : MonadExceptOf Exception CoreM where throw := throw tryCatch := Core.tryCatch -@[inline] def Core.withCatchingRuntimeEx (flag : Bool) (x : CoreM α) : CoreM α := - withReader (fun ctx => { ctx with catchRuntimeEx := flag }) x +class MonadRuntimeException (m : Type → Type) where + tryCatchRuntimeEx (body : m α) (handler : Exception → m α) : m α -@[inline] def mapCoreM [MonadControlT CoreM m] [Monad m] (f : forall {α}, CoreM α → CoreM α) {α} (x : m α) : m α := - controlAt CoreM fun runInBase => f <| runInBase x +export MonadRuntimeException (tryCatchRuntimeEx) -/-- -Execute `x` with `catchRuntimeEx = flag`. That is, given `try x catch ex => h ex`, -if `x` throws a runtime exception, the handler `h` will be invoked if `flag = true` -Recall that --/ -@[inline] def withCatchingRuntimeEx [MonadControlT CoreM m] [Monad m] (x : m α) : m α := - mapCoreM (Core.withCatchingRuntimeEx true) x +instance : MonadRuntimeException CoreM where + tryCatchRuntimeEx := Core.tryCatchRuntimeEx + +@[inline] instance [MonadRuntimeException m] : MonadRuntimeException (ReaderT ρ m) where + tryCatchRuntimeEx := fun x c r => tryCatchRuntimeEx (x r) (fun e => (c e) r) -@[inline] def withoutCatchingRuntimeEx [MonadControlT CoreM m] [Monad m] (x : m α) : m α := - mapCoreM (Core.withCatchingRuntimeEx false) x +@[inline] instance [MonadRuntimeException m] : MonadRuntimeException (StateRefT' ω σ m) where + tryCatchRuntimeEx := fun x c s => tryCatchRuntimeEx (x s) (fun e => c e s) + +@[inline] def mapCoreM [MonadControlT CoreM m] [Monad m] (f : forall {α}, CoreM α → CoreM α) {α} (x : m α) : m α := + controlAt CoreM fun runInBase => f <| runInBase x end Lean diff --git a/src/Lean/Meta/SynthInstance.lean b/src/Lean/Meta/SynthInstance.lean index 6a82fe7c5767..feaace0c8cda 100644 --- a/src/Lean/Meta/SynthInstance.lean +++ b/src/Lean/Meta/SynthInstance.lean @@ -633,12 +633,9 @@ def main (type : Expr) (maxResultSize : Nat) : MetaM (Option AbstractMVarsResult let action : SynthM (Option AbstractMVarsResult) := do newSubgoal (← getMCtx) key mvar Waiter.root synth - -- TODO: it would be nice to have a nice notation for the following idiom - withCatchingRuntimeEx - try - withoutCatchingRuntimeEx do - action.run { maxResultSize := maxResultSize, maxHeartbeats := getMaxHeartbeats (← getOptions) } |>.run' {} - catch ex => + tryCatchRuntimeEx + (action.run { maxResultSize := maxResultSize, maxHeartbeats := getMaxHeartbeats (← getOptions) } |>.run' {}) + fun ex => if ex.isRuntime then throwError "failed to synthesize{indentExpr type}\n{ex.toMessageData}" else diff --git a/src/Lean/Meta/Tactic/Simp/Main.lean b/src/Lean/Meta/Tactic/Simp/Main.lean index a553e22d86c6..2a8bf686e03c 100644 --- a/src/Lean/Meta/Tactic/Simp/Main.lean +++ b/src/Lean/Meta/Tactic/Simp/Main.lean @@ -662,30 +662,30 @@ def main (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Methods := trace[Meta.Tactic.simp.numSteps] "{s.numSteps}" return (r, { s with }) where - simpMain (e : Expr) : SimpM Result := withCatchingRuntimeEx do - try - withoutCatchingRuntimeEx <| simp e - catch ex => - reportDiag (← get).diag - if ex.isRuntime then - throwNestedTacticEx `simp ex - else - throw ex + simpMain (e : Expr) : SimpM Result := + tryCatchRuntimeEx + (simp e) + fun ex => do + reportDiag (← get).diag + if ex.isRuntime then + throwNestedTacticEx `simp ex + else + throw ex def dsimpMain (e : Expr) (ctx : Context) (stats : Stats := {}) (methods : Methods := {}) : MetaM (Expr × Stats) := do withSimpContext ctx do let (r, s) ← dsimpMain e methods.toMethodsRef ctx |>.run { stats with } pure (r, { s with }) where - dsimpMain (e : Expr) : SimpM Expr := withCatchingRuntimeEx do - try - withoutCatchingRuntimeEx <| dsimp e - catch ex => - reportDiag (← get).diag - if ex.isRuntime then - throwNestedTacticEx `simp ex - else - throw ex + dsimpMain (e : Expr) : SimpM Expr := + tryCatchRuntimeEx + (dsimp e) + fun ex => do + reportDiag (← get).diag + if ex.isRuntime then + throwNestedTacticEx `simp ex + else + throw ex end Simp open Simp (SimprocsArray Stats) From a1be9ec8505544801583a41fa3ac1640491958fa Mon Sep 17 00:00:00 2001 From: Kyle Miller Date: Fri, 10 May 2024 15:37:34 -0700 Subject: [PATCH 33/44] chore: tidying up `Lean.unresolveNameGlobal` (#4091) The main loop logic could be simplified, and `if let` could be used to make control flow more obvious. Also adds a check for macro scopes to prevent `unresolveNameGlobal` from returning names with macro scopes in the event there's an alias with one. This is a follow up to #3946. --- src/Lean/ResolveName.lean | 28 ++++++++++++++++++---------- tests/lean/run/2291.lean | 22 ++++++++++++++++++++++ 2 files changed, 40 insertions(+), 10 deletions(-) diff --git a/src/Lean/ResolveName.lean b/src/Lean/ResolveName.lean index 91bef0d99b6e..6c583629b3ae 100644 --- a/src/Lean/ResolveName.lean +++ b/src/Lean/ResolveName.lean @@ -384,6 +384,16 @@ After `open Foo open Boo`, we have def resolveGlobalConstNoOverload [Monad m] [MonadResolveName m] [MonadEnv m] [MonadError m] (id : Syntax) : m Name := do ensureNonAmbiguous id (← resolveGlobalConst id) +/-- +Finds a name that unambiguously resolves to the given name `n₀`. +Considers suffixes of `n₀` and suffixes of aliases of `n₀` when "unresolving". +Aliases are considered first. + +When `fullNames` is true, returns either `n₀` or `_root_.n₀`. + +This function is meant to be used for pretty printing. +If `n₀` is an accessible name, then the result will be an accessible name. +-/ def unresolveNameGlobal [Monad m] [MonadResolveName m] [MonadEnv m] (n₀ : Name) (fullNames := false) : m Name := do if n₀.hasMacroScopes then return n₀ if fullNames then @@ -393,21 +403,19 @@ def unresolveNameGlobal [Monad m] [MonadResolveName m] [MonadEnv m] (n₀ : Name let mut initialNames := (getRevAliases (← getEnv) n₀).toArray initialNames := initialNames.push (rootNamespace ++ n₀) for initialName in initialNames do - match (← unresolveNameCore initialName) with - | none => continue - | some n => return n + if let some n ← unresolveNameCore initialName then + return n return n₀ -- if can't resolve, return the original where unresolveNameCore (n : Name) : m (Option Name) := do + if n.hasMacroScopes then return none let mut revComponents := n.componentsRev let mut candidate := Name.anonymous - for _ in [:revComponents.length] do - match revComponents with - | [] => return none - | cmpt::rest => candidate := Name.appendCore cmpt candidate; revComponents := rest - match (← resolveGlobalName candidate) with - | [(potentialMatch, _)] => if potentialMatch == n₀ then return some candidate else continue - | _ => continue + for cmpt in revComponents do + candidate := Name.appendCore cmpt candidate + if let [(potentialMatch, _)] ← resolveGlobalName candidate then + if potentialMatch == n₀ then + return some candidate return none end Lean diff --git a/tests/lean/run/2291.lean b/tests/lean/run/2291.lean index 0390315bd03c..d7791536e01b 100644 --- a/tests/lean/run/2291.lean +++ b/tests/lean/run/2291.lean @@ -1,3 +1,5 @@ +import Lean.Elab.Command +import Lean.Elab.Open /-! Issue #2291 @@ -6,3 +8,23 @@ The following example would cause the pretty printer to panic. set_option trace.compiler.simp true in #eval [0] + + +/-! +Fixing the above involved changing `Lean.unresolveNameGlobal`. +Here, we also verify that we do not pretty print using any aliases that have macro scopes. +-/ + +open Lean in +elab "add_bad_alias " n:ident : command => withFreshMacroScope do + let declName ← Elab.OpenDecl.resolveNameUsingNamespaces [← getCurrNamespace] n + let badName ← MonadQuotation.addMacroScope `bad + modify fun s => { s with env := addAlias s.env badName declName } + +def f := 1 + +add_bad_alias f + +-- Formerly was info: bad✝ : ℕ +/-- info: f : Nat -/ +#guard_msgs in #check (f) From 25e94f916f1a266aa456aba2b7e570f5aa352403 Mon Sep 17 00:00:00 2001 From: Mac Malone Date: Fri, 10 May 2024 19:04:54 -0400 Subject: [PATCH 34/44] feat: `IO.TaskState` (#4097) Adds `IO.getTaskState` which returns the state of a `Task` in the Lean runtime's task manager. The `TaskState` inductive has 3 constructors: `waiting`, `running`, and `finished`. The `waiting` constructor encompasses the waiting and queued states within the C task object documentation, because the task object does not provide a low cost way to distinguish these different forms of waiting. Furthermore, it seems unlikely for consumers to wish to distinguish between these internal states. The `running` constructor encompasses both the running and promised states in C docs. While not ideal, the C implementation does not provide a way to distinguish between a running `Task` and a waiting `Promise.result` (they both have null closures). --- src/Init/System/IO.lean | 38 ++++++++++++++++++++++++++++++++++- src/include/lean/lean.h | 4 ++-- src/runtime/io.cpp | 4 ++-- src/runtime/object.cpp | 13 ++++++++++-- src/runtime/object.h | 2 +- tests/lean/run/taskState.lean | 25 +++++++++++++++++++++++ 6 files changed, 78 insertions(+), 8 deletions(-) create mode 100644 tests/lean/run/taskState.lean diff --git a/src/Init/System/IO.lean b/src/Init/System/IO.lean index 45eba794fb7d..0c6805cd7b97 100644 --- a/src/Init/System/IO.lean +++ b/src/Init/System/IO.lean @@ -210,8 +210,44 @@ def sleep (ms : UInt32) : BaseIO Unit := /-- Request cooperative cancellation of the task. The task must explicitly call `IO.checkCanceled` to react to the cancellation. -/ @[extern "lean_io_cancel"] opaque cancel : @& Task α → BaseIO Unit +/-- The current state of a `Task` in the Lean runtime's task manager. -/ +inductive TaskState + /-- + The `Task` is waiting to be run. + It can be waiting for dependencies to complete or + sitting in the task manager queue waiting for a thread to run on. + -/ + | waiting + /-- + The `Task` is actively running on a thread or, + in the case of a `Promise`, waiting for a call to `IO.Promise.resolve`. + -/ + | running + /-- + The `Task` has finished running and its result is available. + Calling `Task.get` or `IO.wait` on the task will not block. + -/ + | finished + deriving Inhabited, Repr, DecidableEq, Ord + +instance : LT TaskState := ltOfOrd +instance : LE TaskState := leOfOrd +instance : Min TaskState := minOfLe +instance : Max TaskState := maxOfLe + +protected def TaskState.toString : TaskState → String + | .waiting => "waiting" + | .running => "running" + | .finished => "finished" + +instance : ToString TaskState := ⟨TaskState.toString⟩ + +/-- Returns current state of the `Task` in the Lean runtime's task manager. -/ +@[extern "lean_io_get_task_state"] opaque getTaskState : @& Task α → BaseIO TaskState + /-- Check if the task has finished execution, at which point calling `Task.get` will return immediately. -/ -@[extern "lean_io_has_finished"] opaque hasFinished : @& Task α → BaseIO Bool +@[inline] def hasFinished (task : Task α) : BaseIO Bool := do + return (← getTaskState task) matches .finished /-- Wait for the task to finish, then return its result. -/ @[extern "lean_io_wait"] opaque wait (t : Task α) : BaseIO α := diff --git a/src/include/lean/lean.h b/src/include/lean/lean.h index 80120a38b89d..c189c6235e2d 100644 --- a/src/include/lean/lean.h +++ b/src/include/lean/lean.h @@ -1110,8 +1110,8 @@ static inline lean_obj_res lean_task_get_own(lean_obj_arg t) { LEAN_EXPORT bool lean_io_check_canceled_core(void); /* primitive for implementing `IO.cancel : Task a -> IO Unit` */ LEAN_EXPORT void lean_io_cancel_core(b_lean_obj_arg t); -/* primitive for implementing `IO.hasFinished : Task a -> IO Unit` */ -LEAN_EXPORT bool lean_io_has_finished_core(b_lean_obj_arg t); +/* primitive for implementing `IO.getTaskState : Task a -> IO TaskState` */ +LEAN_EXPORT uint8_t lean_io_get_task_state_core(b_lean_obj_arg t); /* primitive for implementing `IO.waitAny : List (Task a) -> IO (Task a)` */ LEAN_EXPORT b_lean_obj_res lean_io_wait_any_core(b_lean_obj_arg task_list); diff --git a/src/runtime/io.cpp b/src/runtime/io.cpp index 4dddd7afc0b5..fe4a232d9c71 100644 --- a/src/runtime/io.cpp +++ b/src/runtime/io.cpp @@ -1066,8 +1066,8 @@ extern "C" LEAN_EXPORT obj_res lean_io_cancel(b_obj_arg t, obj_arg) { return io_result_mk_ok(box(0)); } -extern "C" LEAN_EXPORT obj_res lean_io_has_finished(b_obj_arg t, obj_arg) { - return io_result_mk_ok(box(lean_io_has_finished_core(t))); +extern "C" LEAN_EXPORT obj_res lean_io_get_task_state(b_obj_arg t, obj_arg) { + return io_result_mk_ok(box(lean_io_get_task_state_core(t))); } extern "C" LEAN_EXPORT obj_res lean_io_wait(obj_arg t, obj_arg) { diff --git a/src/runtime/object.cpp b/src/runtime/object.cpp index 89af9b2604fa..262227eb5e8f 100644 --- a/src/runtime/object.cpp +++ b/src/runtime/object.cpp @@ -1030,8 +1030,17 @@ extern "C" LEAN_EXPORT void lean_io_cancel_core(b_obj_arg t) { g_task_manager->cancel(lean_to_task(t)); } -extern "C" LEAN_EXPORT bool lean_io_has_finished_core(b_obj_arg t) { - return lean_to_task(t)->m_value != nullptr; +extern "C" LEAN_EXPORT uint8_t lean_io_get_task_state_core(b_obj_arg t) { + lean_task_object * o = lean_to_task(t); + if (o->m_imp) { + if (o->m_imp->m_closure) { + return 0; // waiting (waiting/queued) + } else { + return 1; // running (running/promised) + } + } else { + return 2; // finished + } } extern "C" LEAN_EXPORT b_obj_res lean_io_wait_any_core(b_obj_arg task_list) { diff --git a/src/runtime/object.h b/src/runtime/object.h index 04f13fe1b642..f9bf12d6c2da 100644 --- a/src/runtime/object.h +++ b/src/runtime/object.h @@ -288,7 +288,7 @@ inline b_obj_res task_get(b_obj_arg t) { return lean_task_get(t); } inline bool io_check_canceled_core() { return lean_io_check_canceled_core(); } inline void io_cancel_core(b_obj_arg t) { return lean_io_cancel_core(t); } -inline bool io_has_finished_core(b_obj_arg t) { return lean_io_has_finished_core(t); } +inline bool io_get_task_state_core(b_obj_arg t) { return lean_io_get_task_state_core(t); } inline b_obj_res io_wait_any_core(b_obj_arg task_list) { return lean_io_wait_any_core(task_list); } // ======================================= diff --git a/tests/lean/run/taskState.lean b/tests/lean/run/taskState.lean new file mode 100644 index 000000000000..e5e045edec06 --- /dev/null +++ b/tests/lean/run/taskState.lean @@ -0,0 +1,25 @@ +def assertBEq [BEq α] [ToString α] (caption : String) (actual expected : α) : IO Unit := do + unless actual == expected do + throw <| IO.userError <| + s!"{caption}: expected '{expected}', got '{actual}'" + +def test : IO Unit := do + let p1 : IO.Promise Unit ← IO.Promise.new -- resolving queues the task + let p2 : IO.Promise Unit ← IO.Promise.new -- resolved once task is running + let p3 : IO.Promise Unit ← IO.Promise.new -- resolving finishes the task + let t ← BaseIO.mapTask (fun () => do p2.resolve (); IO.wait p3.result) p1.result + assertBEq "p1" (← IO.getTaskState p1.result) .running + assertBEq "p2" (← IO.getTaskState p2.result) .running + assertBEq "p3" (← IO.getTaskState p3.result) .running + assertBEq "t" (← IO.getTaskState t) .waiting + p1.resolve () + assertBEq "p1" (← IO.getTaskState p1.result) .finished + IO.wait p2.result + assertBEq "p2" (← IO.getTaskState p2.result) .finished + assertBEq "t" (← IO.getTaskState t) .running + p3.resolve () + assertBEq "p3" (← IO.getTaskState p3.result) .finished + IO.wait t + assertBEq "t" (← IO.getTaskState t) .finished + +#eval test From a875ae3acf1b7e6e382172b521dcfe27bbfdc57c Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 May 2024 02:07:13 +0200 Subject: [PATCH 35/44] feat: recover from runtime errors in tactics (#4130) closes #3554 --- src/Lean/CoreM.lean | 10 +++++++- src/Lean/Elab/SyntheticMVars.lean | 38 +++++++++++++++---------------- src/Lean/Elab/Tactic/Basic.lean | 18 +++++++-------- tests/lean/run/3554.lean | 37 ++++++++++++++++++++++++++++++ 4 files changed, 74 insertions(+), 29 deletions(-) create mode 100644 tests/lean/run/3554.lean diff --git a/src/Lean/CoreM.lean b/src/Lean/CoreM.lean index fc1923be0627..2aa865b78cff 100644 --- a/src/Lean/CoreM.lean +++ b/src/Lean/CoreM.lean @@ -244,8 +244,16 @@ protected def withIncRecDepth [Monad m] [MonadControlT CoreM m] (x : m α) : m -- should never be visible to users! throw <| Exception.error .missing "elaboration interrupted" +register_builtin_option debug.moduleNameAtTimeout : Bool := { + defValue := true + group := "debug" + descr := "include module name in deterministic timeout error messages.\nRemark: we set this option to false to increase the stability of our test suite" +} + def throwMaxHeartbeat (moduleName : Name) (optionName : Name) (max : Nat) : CoreM Unit := do - let msg := s!"(deterministic) timeout at `{moduleName}`, maximum number of heartbeats ({max/1000}) has been reached\nuse `set_option {optionName} ` to set the limit\nuse `set_option {diagnostics.name} true` to get diagnostic information" + let includeModuleName := debug.moduleNameAtTimeout.get (← getOptions) + let atModuleName := if includeModuleName then s!" at `{moduleName}`" else "" + let msg := s!"(deterministic) timeout{atModuleName}, maximum number of heartbeats ({max/1000}) has been reached\nuse `set_option {optionName} ` to set the limit\nuse `set_option {diagnostics.name} true` to get diagnostic information" throw <| Exception.error (← getRef) (MessageData.ofFormat (Std.Format.text msg)) def checkMaxHeartbeatsCore (moduleName : String) (optionName : Name) (max : Nat) : CoreM Unit := do diff --git a/src/Lean/Elab/SyntheticMVars.lean b/src/Lean/Elab/SyntheticMVars.lean index f605436c94c6..bafa9ec11407 100644 --- a/src/Lean/Elab/SyntheticMVars.lean +++ b/src/Lean/Elab/SyntheticMVars.lean @@ -340,26 +340,26 @@ mutual Regarding issue #1380, we addressed the issue by avoiding the elaboration postponement step. However, the same issue can happen in more complicated scenarios. -/ - try - let remainingGoals ← withInfoHole mvarId <| Tactic.run mvarId do - withTacticInfoContext tacticCode do - -- also put an info node on the `by` keyword specifically -- the token may be `canonical` and thus shown in the info - -- view even though it is synthetic while a node like `tacticCode` never is (#1990) - withTacticInfoContext tacticCode[0] do - evalTactic code - synthesizeSyntheticMVars (postpone := .no) - unless remainingGoals.isEmpty do - if report then - reportUnsolvedGoals remainingGoals + tryCatchRuntimeEx + (do let remainingGoals ← withInfoHole mvarId <| Tactic.run mvarId do + withTacticInfoContext tacticCode do + -- also put an info node on the `by` keyword specifically -- the token may be `canonical` and thus shown in the info + -- view even though it is synthetic while a node like `tacticCode` never is (#1990) + withTacticInfoContext tacticCode[0] do + evalTactic code + synthesizeSyntheticMVars (postpone := .no) + unless remainingGoals.isEmpty do + if report then + reportUnsolvedGoals remainingGoals + else + throwError "unsolved goals\n{goalsToMessageData remainingGoals}") + fun ex => do + if report && (← read).errToSorry then + for mvarId in (← getMVars (mkMVar mvarId)) do + mvarId.admit + logException ex else - throwError "unsolved goals\n{goalsToMessageData remainingGoals}" - catch ex => - if report && (← read).errToSorry then - for mvarId in (← getMVars (mkMVar mvarId)) do - mvarId.admit - logException ex - else - throw ex + throw ex /-- Try to synthesize the given pending synthetic metavariable. -/ private partial def synthesizeSyntheticMVar (mvarId : MVarId) (postponeOnError : Bool) (runTactics : Bool) : TermElabM Bool := do diff --git a/src/Lean/Elab/Tactic/Basic.lean b/src/Lean/Elab/Tactic/Basic.lean index 70064dfcdbf3..6c6610bf4703 100644 --- a/src/Lean/Elab/Tactic/Basic.lean +++ b/src/Lean/Elab/Tactic/Basic.lean @@ -231,15 +231,15 @@ def closeUsingOrAdmit (tac : TacticM Unit) : TacticM Unit := do /- Important: we must define `closeUsingOrAdmit` before we define the instance `MonadExcept` for `TacticM` since it backtracks the state including error messages. -/ let mvarId :: mvarIds ← getUnsolvedGoals | throwNoGoalsToBeSolved - try - focusAndDone tac - catch ex => - if (← read).recover then - logException ex - admitGoal mvarId - setGoals mvarIds - else - throw ex + tryCatchRuntimeEx + (focusAndDone tac) + fun ex => do + if (← read).recover then + logException ex + admitGoal mvarId + setGoals mvarIds + else + throw ex instance : MonadBacktrack SavedState TacticM where saveState := Tactic.saveState diff --git a/tests/lean/run/3554.lean b/tests/lean/run/3554.lean new file mode 100644 index 000000000000..ac3afc8ef353 --- /dev/null +++ b/tests/lean/run/3554.lean @@ -0,0 +1,37 @@ +def foo : Nat → Nat +| 0 => 0 +| n+1 => foo n + 1 + +set_option debug.moduleNameAtTimeout false +/-- +error: (deterministic) timeout, maximum number of heartbeats (100) has been reached +use `set_option maxHeartbeats ` to set the limit +use `set_option diagnostics true` to get diagnostic information +-/ +#guard_msgs in +set_option maxHeartbeats 100 in +theorem bar : True := by + simp [show foo 1000 = 1000 from rfl] + +/-- +info: theorem bar : True := +sorryAx True true +-/ +#guard_msgs in +#print bar + +/-- +error: maximum recursion depth has been reached +use `set_option maxRecDepth ` to increase limit +use `set_option diagnostics true` to get diagnostic information +-/ +#guard_msgs in +theorem bar2 : True := by + simp [show foo 1000 = 1000 from rfl] + +/-- +info: theorem bar2 : True := +sorryAx True true +-/ +#guard_msgs in +#print bar2 From 147aeaea456334ea9c133ea72d6fb01f666b906a Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 May 2024 02:47:33 +0200 Subject: [PATCH 36/44] test: for issue 2558 (#4133) Issue has been fixed by another PR. closes #2558 --- tests/lean/run/2558.lean | 3 +++ 1 file changed, 3 insertions(+) create mode 100644 tests/lean/run/2558.lean diff --git a/tests/lean/run/2558.lean b/tests/lean/run/2558.lean new file mode 100644 index 000000000000..948c1bab869c --- /dev/null +++ b/tests/lean/run/2558.lean @@ -0,0 +1,3 @@ +def even2 : sorry := + have e : 1 + 1 = 2 := Eq.refl _ + sorry From 7db8e6482e1dc46fd1070b3bf549112f02a4c05e Mon Sep 17 00:00:00 2001 From: Leonardo de Moura Date: Sat, 11 May 2024 04:10:40 +0200 Subject: [PATCH 37/44] fix: auto/option params should not break `sorry` (#4132) closes #2649 --- src/Lean/Elab/BuiltinNotation.lean | 2 +- tests/lean/run/2649.lean | 3 +++ 2 files changed, 4 insertions(+), 1 deletion(-) create mode 100644 tests/lean/run/2649.lean diff --git a/src/Lean/Elab/BuiltinNotation.lean b/src/Lean/Elab/BuiltinNotation.lean index cca120f25ca0..65b12af63902 100644 --- a/src/Lean/Elab/BuiltinNotation.lean +++ b/src/Lean/Elab/BuiltinNotation.lean @@ -205,7 +205,7 @@ private def elabTParserMacroAux (prec lhsPrec e : Term) : TermElabM Syntax := do | _ => Macro.throwUnsupported @[builtin_term_elab «sorry»] def elabSorry : TermElab := fun stx expectedType? => do - let stxNew ← `(sorryAx _ false) + let stxNew ← `(@sorryAx _ false) -- Remark: we use `@` to ensure `sorryAx` will not consume auot params withMacroExpansion stx stxNew <| elabTerm stxNew expectedType? /-- Return syntax `Prod.mk elems[0] (Prod.mk elems[1] ... (Prod.mk elems[elems.size - 2] elems[elems.size - 1])))` -/ diff --git a/tests/lean/run/2649.lean b/tests/lean/run/2649.lean new file mode 100644 index 000000000000..9dbd36888e8d --- /dev/null +++ b/tests/lean/run/2649.lean @@ -0,0 +1,3 @@ +def foo1 : (_ : Nat := 0) → Nat := sorry + +def foo2 : (_ : Nat := by exact 0) → Nat := sorry From 0d9af1b7777b5671a7eb061fc1f079ca20b6144d Mon Sep 17 00:00:00 2001 From: Marcus Rossel Date: Sun, 12 May 2024 18:45:56 +0200 Subject: [PATCH 38/44] fix: typo in `Meta.unfoldProjInstWhenInstances?` (#4139) --- src/Lean/Meta/ExprDefEq.lean | 4 ++-- src/Lean/Meta/WHNF.lean | 6 +++--- 2 files changed, 5 insertions(+), 5 deletions(-) diff --git a/src/Lean/Meta/ExprDefEq.lean b/src/Lean/Meta/ExprDefEq.lean index 1876518f2466..9cd8b3b6c2b9 100644 --- a/src/Lean/Meta/ExprDefEq.lean +++ b/src/Lean/Meta/ExprDefEq.lean @@ -1904,8 +1904,8 @@ private def isDefEqUnitLike (t : Expr) (s : Expr) : MetaM Bool := do -/ private def isDefEqProjInst (t : Expr) (s : Expr) : MetaM LBool := do if (← getTransparency) != .instances then return .undef - let t? ← unfoldProjInstWhenIntances? t - let s? ← unfoldProjInstWhenIntances? s + let t? ← unfoldProjInstWhenInstances? t + let s? ← unfoldProjInstWhenInstances? s if t?.isSome || s?.isSome then toLBoolM <| Meta.isExprDefEqAux (t?.getD t) (s?.getD s) else diff --git a/src/Lean/Meta/WHNF.lean b/src/Lean/Meta/WHNF.lean index ee84d98b492c..ebd379279b36 100644 --- a/src/Lean/Meta/WHNF.lean +++ b/src/Lean/Meta/WHNF.lean @@ -712,11 +712,11 @@ mutual | _ => return none /-- - Auxiliary method for unfolding a class projection. when transparency is set to `TransparencyMode.instances`. + Auxiliary method for unfolding a class projection when transparency is set to `TransparencyMode.instances`. Recall that class instance projections are not marked with `[reducible]` because we want them to be in "reducible canonical form". -/ - partial def unfoldProjInstWhenIntances? (e : Expr) : MetaM (Option Expr) := do + partial def unfoldProjInstWhenInstances? (e : Expr) : MetaM (Option Expr) := do if (← getTransparency) != TransparencyMode.instances then return none else @@ -726,7 +726,7 @@ mutual partial def unfoldDefinition? (e : Expr) : MetaM (Option Expr) := match e with | .app f _ => - matchConstAux f.getAppFn (fun _ => unfoldProjInstWhenIntances? e) fun fInfo fLvls => do + matchConstAux f.getAppFn (fun _ => unfoldProjInstWhenInstances? e) fun fInfo fLvls => do if fInfo.levelParams.length != fLvls.length then return none else From b8f2f28e0dc93dff65c5715bc041804d49560f59 Mon Sep 17 00:00:00 2001 From: Joachim Breitner Date: Sun, 12 May 2024 22:39:14 +0200 Subject: [PATCH 39/44] fix: check that funind-reserved names are available (#4135) I did not introduce `inductTheoremSuffix` etc, it seems more direct to just spell out the suffix here. If we ever change it there are many occurrences where they need to be changed anyways, so the definition doesn't seem to save much work or add that much robustness. --- src/Lean/Elab/PreDefinition/Main.lean | 6 ++ tests/lean/funind_reserved.lean | 96 ++++++++++++++++++++ tests/lean/funind_reserved.lean.expected.out | 6 ++ 3 files changed, 108 insertions(+) create mode 100644 tests/lean/funind_reserved.lean create mode 100644 tests/lean/funind_reserved.lean.expected.out diff --git a/src/Lean/Elab/PreDefinition/Main.lean b/src/Lean/Elab/PreDefinition/Main.lean index 200fa29f5cbf..d65b178aafd9 100644 --- a/src/Lean/Elab/PreDefinition/Main.lean +++ b/src/Lean/Elab/PreDefinition/Main.lean @@ -90,6 +90,11 @@ private def addAsAxioms (preDefs : Array PreDefinition) : TermElabM Unit := do applyAttributesOf #[preDef] AttributeApplicationTime.afterTypeChecking applyAttributesOf #[preDef] AttributeApplicationTime.afterCompilation +def ensureFunIndReservedNamesAvailable (preDefs : Array PreDefinition) : MetaM Unit := do + preDefs.forM fun preDef => + withRef preDef.ref <| ensureReservedNameAvailable preDef.declName "induct" + withRef preDefs[0]!.ref <| ensureReservedNameAvailable preDefs[0]!.declName "mutual_induct" + def addPreDefinitions (preDefs : Array PreDefinition) : TermElabM Unit := withLCtx {} {} do for preDef in preDefs do trace[Elab.definition.body] "{preDef.declName} : {preDef.type} :=\n{preDef.value}" @@ -121,6 +126,7 @@ def addPreDefinitions (preDefs : Array PreDefinition) : TermElabM Unit := withLC addAndCompilePartial preDefs preDefs.forM (·.termination.ensureNone "partial") else + ensureFunIndReservedNamesAvailable preDefs try let hasHints := preDefs.any fun preDef => preDef.termination.isNotNone if hasHints then diff --git a/tests/lean/funind_reserved.lean b/tests/lean/funind_reserved.lean new file mode 100644 index 000000000000..d822a60ef70c --- /dev/null +++ b/tests/lean/funind_reserved.lean @@ -0,0 +1,96 @@ + +-- Test that the reserved name availability is checked at function definition time + +namespace Nonrec + +def foo.induct := 1 + +def foo : (n : Nat) → Nat -- no error (yet) + | 0 => 0 + | n+1 => n + +end Nonrec + +namespace Structural + +def foo.induct := 1 + +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => foo n + +end Structural + +namespace WF + +def foo.induct := 1 + +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => foo n +termination_by n => n + +end WF + +namespace Mutual1 + +def foo.induct := 1 + +mutual +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => bar n +termination_by n => n + +def bar : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => foo n +termination_by n => n +end + +end Mutual1 + +namespace Mutual2 + +def bar.induct := 1 + +mutual +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => bar n +termination_by n => n + +def bar : (n : Nat) → Nat + | 0 => 0 + | n+1 => foo n +termination_by n => n +end + +end Mutual2 + +namespace Mutual3 + +def foo.mutual_induct := 1 + +mutual +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => bar n +termination_by n => n + +def bar : (n : Nat) → Nat + | 0 => 0 + | n+1 => foo n +termination_by n => n +end + +end Mutual3 + +namespace Nested + +def foo : (n : Nat) → Nat -- error + | 0 => 0 + | n+1 => foo n + where induct := 1 + +end Nested diff --git a/tests/lean/funind_reserved.lean.expected.out b/tests/lean/funind_reserved.lean.expected.out new file mode 100644 index 000000000000..36e3a5ca5217 --- /dev/null +++ b/tests/lean/funind_reserved.lean.expected.out @@ -0,0 +1,6 @@ +funind_reserved.lean:18:4-18:7: error: failed to declare `Structural.foo` because `Structural.foo.induct` has already been declared +funind_reserved.lean:28:4-28:7: error: failed to declare `WF.foo` because `WF.foo.induct` has already been declared +funind_reserved.lean:40:4-40:7: error: failed to declare `Mutual1.foo` because `Mutual1.foo.induct` has already been declared +funind_reserved.lean:63:4-63:7: error: failed to declare `Mutual2.bar` because `Mutual2.bar.induct` has already been declared +funind_reserved.lean:76:4-76:7: error: failed to declare `Mutual3.foo` because `Mutual3.foo.mutual_induct` has already been declared +funind_reserved.lean:91:4-91:7: error: failed to declare `Nested.foo` because `Nested.foo.induct` has already been declared From f74980ccee82ca2abdae65dcbc5571d4640ed076 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 13 May 2024 09:06:48 +1000 Subject: [PATCH 40/44] chore: incorrect lemma resolution in omega (#4141) Fixes #4138. --- src/Lean/Elab/Tactic/Omega/OmegaM.lean | 2 +- tests/lean/run/omega.lean | 4 ++++ 2 files changed, 5 insertions(+), 1 deletion(-) diff --git a/src/Lean/Elab/Tactic/Omega/OmegaM.lean b/src/Lean/Elab/Tactic/Omega/OmegaM.lean index 57048157b75a..091843360645 100644 --- a/src/Lean/Elab/Tactic/Omega/OmegaM.lean +++ b/src/Lean/Elab/Tactic/Omega/OmegaM.lean @@ -199,7 +199,7 @@ def analyzeAtom (e : Expr) : OmegaM (HashSet Expr) := do | some _ => let b_pos := mkApp4 (.const ``LT.lt [0]) (.const ``Int []) (.const ``Int.instLTInt []) (toExpr (0 : Int)) b - let pow_pos := mkApp3 (.const ``Int.pos_pow_of_pos []) b exp (← mkDecideProof b_pos) + let pow_pos := mkApp3 (.const ``Lean.Omega.Int.pos_pow_of_pos []) b exp (← mkDecideProof b_pos) pure <| HashSet.empty.insert (mkApp3 (.const ``Int.emod_nonneg []) x k (mkApp3 (.const ``Int.ne_of_gt []) k (toExpr (0 : Int)) pow_pos)) |>.insert diff --git a/tests/lean/run/omega.lean b/tests/lean/run/omega.lean index 5c6b59331fe3..262ecfe2e06e 100644 --- a/tests/lean/run/omega.lean +++ b/tests/lean/run/omega.lean @@ -426,6 +426,10 @@ example (x e : Nat) (hx : x < 2^(e.succ)) : x < 2^e * 2 := by omega -- Check that this works for integer base. example (x : Int) (e : Nat) (hx : x < (2 : Int)^(e+1)) : x < 2^e * 2 := by omega +example (n : Nat) (i : Int) (h2n : (2 : Int) ^ n = ↑((2 : Nat) ^ (n : Nat))) + (hlt : i % 2 ^ n < 2 ^ n) : 2 ^ n ≠ 0 := by + omega + /-! ### Ground terms -/ example : 2^7 < 165 := by omega From 799923d145960c6f2515b510111df2c170b34ca3 Mon Sep 17 00:00:00 2001 From: Kim Morrison Date: Mon, 13 May 2024 16:18:44 +1000 Subject: [PATCH 41/44] chore: move have to decreasing_by in substrEq.loop (#4143) Currently this causes linter warnings downstream in proofs that unfold substrEq.loop. --- src/Init/Data/String/Basic.lean | 6 ++++-- 1 file changed, 4 insertions(+), 2 deletions(-) diff --git a/src/Init/Data/String/Basic.lean b/src/Init/Data/String/Basic.lean index 1ea50f2e6869..e49870e9ec51 100644 --- a/src/Init/Data/String/Basic.lean +++ b/src/Init/Data/String/Basic.lean @@ -683,13 +683,15 @@ def substrEq (s1 : String) (off1 : String.Pos) (s2 : String) (off2 : String.Pos) off1.byteIdx + sz ≤ s1.endPos.byteIdx && off2.byteIdx + sz ≤ s2.endPos.byteIdx && loop off1 off2 { byteIdx := off1.byteIdx + sz } where loop (off1 off2 stop1 : Pos) := - if h : off1.byteIdx < stop1.byteIdx then + if _h : off1.byteIdx < stop1.byteIdx then let c₁ := s1.get off1 let c₂ := s2.get off2 - have := Nat.sub_lt_sub_left h (Nat.add_lt_add_left (one_le_csize c₁) off1.1) c₁ == c₂ && loop (off1 + c₁) (off2 + c₂) stop1 else true termination_by stop1.1 - off1.1 + decreasing_by + have := Nat.sub_lt_sub_left _h (Nat.add_lt_add_left (one_le_csize c₁) off1.1) + decreasing_tactic /-- Return true iff `p` is a prefix of `s` -/ def isPrefixOf (p : String) (s : String) : Bool := From a17c3f424cb5f2390e645d8a2cf84e69a1b96a37 Mon Sep 17 00:00:00 2001 From: Siddharth Date: Mon, 13 May 2024 13:35:59 +0100 Subject: [PATCH 42/44] feat: BitVec.shiftLeft_shiftLeft, BitVec.shiftRight_shiftRight (#4148) Closes two `sorry`s at https://github.com/leanprover/leansat/pull/64/files. --------- Co-authored-by: Kim Morrison --- src/Init/Data/BitVec/Lemmas.lean | 16 ++++++++++++++++ 1 file changed, 16 insertions(+) diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index a998800ea12a..13f4bf8a8a95 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -608,6 +608,17 @@ theorem shiftLeftZeroExtend_eq {x : BitVec w} : (shiftLeftZeroExtend x i).msb = x.msb := by simp [shiftLeftZeroExtend_eq, BitVec.msb] +theorem BitVec.shiftLeft_shiftLeft {w : Nat} (x : BitVec w) (n m : Nat) : + (x <<< n) <<< m = x <<< (n + m) := by + ext i + simp only [getLsb_shiftLeft, Fin.is_lt, decide_True, Bool.true_and] + rw [show i - (n + m) = (i - m - n) by omega] + cases h₂ : decide (i < m) <;> + cases h₃ : decide (i - m < w) <;> + cases h₄ : decide (i - m < n) <;> + cases h₅ : decide (i < n + m) <;> + simp at * <;> omega + /-! ### ushiftRight -/ @[simp, bv_toNat] theorem toNat_ushiftRight (x : BitVec n) (i : Nat) : @@ -693,6 +704,11 @@ theorem msb_append {x : BitVec w} {y : BitVec v} : simp only [getLsb_append, cond_eq_if] split <;> simp [*] +theorem BitVec.shiftRight_shiftRight (w : Nat) (x : BitVec w) (n m : Nat) : + (x >>> n) >>> m = x >>> (n + m) := by + ext i + simp [Nat.add_assoc n m i] + /-! ### rev -/ theorem getLsb_rev (x : BitVec w) (i : Fin w) : From e39d27c8a71e969a2d2379e17084110271b3b393 Mon Sep 17 00:00:00 2001 From: Alex Keizer Date: Mon, 13 May 2024 18:03:38 +0100 Subject: [PATCH 43/44] WIP --- src/Init/Data/BitVec/Bitblast.lean | 99 ++++++++++++++++++++++++++++++ 1 file changed, 99 insertions(+) diff --git a/src/Init/Data/BitVec/Bitblast.lean b/src/Init/Data/BitVec/Bitblast.lean index ee2e66d7922a..a9f7463b2e02 100644 --- a/src/Init/Data/BitVec/Bitblast.lean +++ b/src/Init/Data/BitVec/Bitblast.lean @@ -184,4 +184,103 @@ theorem bit_neg_eq_neg (x : BitVec w) : -x = (adc (((iunfoldr (fun (i : Fin w) c simp [← sub_toAdd, BitVec.sub_add_cancel] · simp [bit_not_testBit x _] +/-! +### Multiplication +We implement [Booth's multiplication circuit](https://en.wikipedia.org/wiki/Booth%27s_multiplication_algorithm) +on bitvectors, and show that this circuit is equal to our straightforward `BitVec.mul` implementation. +-/ + +def mulAdd (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) : BitVec (w+v) := + let x : BitVec (w+v) := x.zeroExtend' (le_add_right w v) + Prod.snd <| iunfoldr (s:=a) fun (i : Fin (w+v)) a => + let a := if y.getLsb i = true then a + x else a + (a >>> 1, a.getLsb 0) + +def mulAddAccumulator (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) (i : Nat) : BitVec (w+v) := + (a + (x.zeroExtend' <| le_add_right w v) * ((y.extractLsb' 0 i).zeroExtend _) ) >>> i + +@[simp] +theorem zeroExtend_zero_width (x : BitVec 0) : zeroExtend w x = 0#w := by + sorry + +@[simp] theorem shiftRight_zero (x : BitVec w) : x >>> 0 = x := rfl +@[simp] theorem mul_zero (x : BitVec w) : x * 0#w = 0#w := rfl + +theorem extractLsb'_succ_eq_concat (x : BitVec w) (s n : Nat) : + x.extractLsb' s (n+1) = cons (x.getLsb (s+n)) (x.extractLsb' s n) := by + sorry + +theorem mulAdd_spec (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) : + mulAdd a x y = a + (x.zeroExtend' <| le_add_right w v) * (y.zeroExtend' <| le_add_left v w) := by + simp only [mulAdd] + rw [iunfoldr_replace (state := mulAddAccumulator a x y)] + · simp [mulAddAccumulator] + · intro i + simp only [mulAddAccumulator, Prod.mk.injEq] + simp only [extractLsb'_succ_eq_concat y 0 i, Nat.zero_add] + cases y.getLsb i <;> simp + · sorry + · sorry + +@[simp] theorem zeroExtend'_mul_zeroExtend' (x y : BitVec w) (h : w ≤ v) : + x.zeroExtend' h * y.zeroExtend' h = (x * y).zeroExtend' h := by + sorry + + +@[simp] theorem zeroExtend'_rfl (x : BitVec w) (h : w ≤ w := by rfl) : x.zeroExtend' h = x := rfl + +@[simp] +theorem truncate_zeroExtend' (x : BitVec w) (h : w ≤ v) : truncate w (x.zeroExtend' h) = x := by + simp [truncate, zeroExtend] + intro h' + have h_eq : w = v := Nat.le_antisymm h h' + subst h_eq + simp [h'] + +theorem mul_eq_mulAdd (x y : BitVec w) : + x * y = (mulAdd 0 x y).truncate w := by + simp [mulAdd_spec] + +@[simp] +theorem extractLsb'_zero (x : BitVec w) : extractLsb' 0 n x = truncate n x := by + simp [extractLsb'] + +@[simp] +theorem extractLsb'_succ_concat : extractLsb' (start+1) n (concat x a) = extractLsb' start n x := by + simp [extractLsb'] + sorry + +-- theorem mulAdd_eq + +theorem mul_eq_mulAdd (x y : BitVec w) : + x * y = (mulAdd 0 x y).truncate _ := by + suffices ∀ {v w} (x : BitVec (w+v)) (y : BitVec w) (z : BitVec v), + x * (y ++ z) = (mulAdd (x*z) x y).truncate _ + by simpa using @this 0 w x y 0 + induction w + case zero => + sorry + case succ w ih => + have ⟨x, x₀, hx⟩ : ∃ (x' : BitVec w) (x₀ : Bool), x = BitVec.concat x' x₀ := sorry + have ⟨y, y₀, hy⟩ : ∃ (y' : BitVec w) (y₀ : Bool), y = BitVec.concat y' y₀ := sorry + subst hx hy + cases y₀ <;> simp [mulAdd] + · simp [extractLsb'] + rw [show 0#w = 0 from rfl, ← ih] + · sorry + +def mulC (x y : BitVec w) : BitVec w := + go _ + where go (acc : BitVec w) (x y : BitVec w) : BitVec w + +-- def boothMul (x y : BitVec w) : BitVec w := +-- let a : BitVec (w+w+1) := x ++ (0 : BitVec (w+1)) +-- let s : BitVec (w+w+1) := -x ++ (0 : BitVec (w+1)) +-- let p : BitVec (w+w+1) := (0 : BitVec w) ++ (y : BitVec w) ++ (0 : BitVec 1) +-- go a s p w +-- where +-- go (a s p : BitVec (w+w+1)) : Nat → BitVec w +-- | 0 => p +-- | n+1 => + end BitVec From 738db83af8a2c3d841c4fecd36b8d42787de6d6f Mon Sep 17 00:00:00 2001 From: Alex Keizer Date: Tue, 14 May 2024 16:22:29 +0100 Subject: [PATCH 44/44] WIP --- src/Init/Data/BitVec/Bitblast.lean | 139 +++++++++++++++++------------ src/Init/Data/BitVec/Lemmas.lean | 2 +- 2 files changed, 82 insertions(+), 59 deletions(-) diff --git a/src/Init/Data/BitVec/Bitblast.lean b/src/Init/Data/BitVec/Bitblast.lean index a9f7463b2e02..2a5e03430a10 100644 --- a/src/Init/Data/BitVec/Bitblast.lean +++ b/src/Init/Data/BitVec/Bitblast.lean @@ -190,47 +190,112 @@ We implement [Booth's multiplication circuit](https://en.wikipedia.org/wiki/Boot on bitvectors, and show that this circuit is equal to our straightforward `BitVec.mul` implementation. -/ -def mulAdd (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) : BitVec (w+v) := - let x : BitVec (w+v) := x.zeroExtend' (le_add_right w v) - Prod.snd <| iunfoldr (s:=a) fun (i : Fin (w+v)) a => +def mulAdd (a x y : BitVec w) : BitVec w := + Prod.snd <| iunfoldr (s:=a) fun (i : Fin w) a => let a := if y.getLsb i = true then a + x else a (a >>> 1, a.getLsb 0) -def mulAddAccumulator (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) (i : Nat) : BitVec (w+v) := - (a + (x.zeroExtend' <| le_add_right w v) * ((y.extractLsb' 0 i).zeroExtend _) ) >>> i +def mulAddAccumulator (a x y : BitVec w) (i : Nat) : BitVec w := + (a + x * (y.truncate i |>.zeroExtend _)) >>> i + +@[simp] theorem truncate_zero : truncate 0 x = 0#0 := of_length_zero + +@[simp] theorem mul_zero (x : BitVec w) : x * 0#w = 0#w := rfl +@[simp] theorem shiftRight_zero (x : BitVec w) : x >>> 0 = x := rfl +@[simp] theorem shiftLeft_zero (x : BitVec w) : x <<< 0 = x := by apply eq_of_toNat_eq; simp + +theorem mulAddAccumulator_zero (a x y : BitVec w) : mulAddAccumulator a x y 0 = a := by + simp [mulAddAccumulator] + +theorem Nat.shiftRight_add' (m n k : Nat) : + m >>> n + k = (m + (k <<< n)) >>> n := by + sorry + +theorem shiftRight_add' (x y : BitVec w) (n : Nat) : + x >>> n + y = (x + (y <<< n)) >>> n := by + sorry + +#check BitVec.shiftRight_shiftRight + +theorem zeroExtend_truncate_eq_and (x : BitVec w) (i : Nat) : + zeroExtend w (x.truncate i) = x &&& ((-1 : BitVec _) >>> (w-i)) := by + sorry + +theorem add_shiftRight (x y : BitVec w) (n : Nat) : (x + y) >>> n = (x >>> n) + (y >>> n) := by + sorry + +@[simp] theorem zero_shiftRight (w n : Nat) : 0#w >>> n = 0#w := by + sorry + +theorem mod_two_pow_shiftRight (x m n : Nat) : (x % 2^m) >>> n = (x >>> n) % (2^(m+n)) := by + induction n + case zero => rfl + case succ n ih => + simp [shiftRight_succ] + sorry + +theorem shiftLeft_shiftRight_eq_zeroExtend_truncate (x : BitVec w) (i : Nat) : + x <<< i >>> i = zeroExtend w (truncate (w-i) x) := by + apply eq_of_toNat_eq + simp only [toNat_ushiftRight, toNat_shiftLeft, toNat_truncate] + induction i + case a.zero => simp + case a.succ i ih => + rw [mod_two_pow_shiftRight] + sorry + +theorem mulAddAccumulator_succ (a x y : BitVec w) : + mulAddAccumulator a x y (i+1) + = (mulAddAccumulator a x y i >>> 1) + + bif y.getLsb (i+1) then (x.truncate (i+1) |>.zeroExtend _) else 0#w := by + -- ext j + simp only [mulAddAccumulator, natCast_eq_ofNat, BitVec.shiftRight_shiftRight] + have : + x * zeroExtend w (truncate (i + 1) y) + = x * zeroExtend w (truncate i y) + (bif y.getLsb (i+1) then x <<< (i+1) else 0) := by + simp [← shiftLeft_shiftRight_eq_zeroExtend_truncate] + rw [this, ← BitVec.add_assoc, add_shiftRight] + congr + cases y.getLsb (i+1) + · simp + · simp; sorry + + + @[simp] theorem zeroExtend_zero_width (x : BitVec 0) : zeroExtend w x = 0#w := by sorry -@[simp] theorem shiftRight_zero (x : BitVec w) : x >>> 0 = x := rfl -@[simp] theorem mul_zero (x : BitVec w) : x * 0#w = 0#w := rfl +-- @[simp] theorem shiftRight_zero (x : BitVec w) : x >>> 0 = x := rfl +-- @[simp] theorem mul_zero (x : BitVec w) : x * 0#w = 0#w := rfl theorem extractLsb'_succ_eq_concat (x : BitVec w) (s n : Nat) : x.extractLsb' s (n+1) = cons (x.getLsb (s+n)) (x.extractLsb' s n) := by sorry -theorem mulAdd_spec (a : BitVec (w+v)) (x : BitVec w) (y : BitVec v) : - mulAdd a x y = a + (x.zeroExtend' <| le_add_right w v) * (y.zeroExtend' <| le_add_left v w) := by +theorem mulAdd_spec (a x y : BitVec w): + mulAdd a x y = a + x * y := by simp only [mulAdd] rw [iunfoldr_replace (state := mulAddAccumulator a x y)] - · simp [mulAddAccumulator] + · simp [mulAddAccumulator, Nat.mod_one] · intro i - simp only [mulAddAccumulator, Prod.mk.injEq] - simp only [extractLsb'_succ_eq_concat y 0 i, Nat.zero_add] + simp only [mulAddAccumulator, Prod.mk.injEq, natCast_eq_ofNat] cases y.getLsb i <;> simp · sorry · sorry -@[simp] theorem zeroExtend'_mul_zeroExtend' (x y : BitVec w) (h : w ≤ v) : - x.zeroExtend' h * y.zeroExtend' h = (x * y).zeroExtend' h := by +theorem getLsb_mul (x y : BitVec w) (i : Fin w) : + (x * y).getLsb i = Bool.xor (x.getLsb i && y.getLsb i) ((mulAddAccumulator 0 x y i).getLsb 0) := by sorry +theorem zeroExtend'_mul_zeroExtend' (x y : BitVec w) (h : w ≤ v) : + x.zeroExtend' h * y.zeroExtend' h = (x * y).zeroExtend' h := by + sorry @[simp] theorem zeroExtend'_rfl (x : BitVec w) (h : w ≤ w := by rfl) : x.zeroExtend' h = x := rfl -@[simp] -theorem truncate_zeroExtend' (x : BitVec w) (h : w ≤ v) : truncate w (x.zeroExtend' h) = x := by +@[simp] theorem truncate_zeroExtend' (x : BitVec w) (h : w ≤ v) : truncate w (x.zeroExtend' h) = x := by simp [truncate, zeroExtend] intro h' have h_eq : w = v := Nat.le_antisymm h h' @@ -241,46 +306,4 @@ theorem mul_eq_mulAdd (x y : BitVec w) : x * y = (mulAdd 0 x y).truncate w := by simp [mulAdd_spec] -@[simp] -theorem extractLsb'_zero (x : BitVec w) : extractLsb' 0 n x = truncate n x := by - simp [extractLsb'] - -@[simp] -theorem extractLsb'_succ_concat : extractLsb' (start+1) n (concat x a) = extractLsb' start n x := by - simp [extractLsb'] - sorry - --- theorem mulAdd_eq - -theorem mul_eq_mulAdd (x y : BitVec w) : - x * y = (mulAdd 0 x y).truncate _ := by - suffices ∀ {v w} (x : BitVec (w+v)) (y : BitVec w) (z : BitVec v), - x * (y ++ z) = (mulAdd (x*z) x y).truncate _ - by simpa using @this 0 w x y 0 - induction w - case zero => - sorry - case succ w ih => - have ⟨x, x₀, hx⟩ : ∃ (x' : BitVec w) (x₀ : Bool), x = BitVec.concat x' x₀ := sorry - have ⟨y, y₀, hy⟩ : ∃ (y' : BitVec w) (y₀ : Bool), y = BitVec.concat y' y₀ := sorry - subst hx hy - cases y₀ <;> simp [mulAdd] - · simp [extractLsb'] - rw [show 0#w = 0 from rfl, ← ih] - · sorry - -def mulC (x y : BitVec w) : BitVec w := - go _ - where go (acc : BitVec w) (x y : BitVec w) : BitVec w - --- def boothMul (x y : BitVec w) : BitVec w := --- let a : BitVec (w+w+1) := x ++ (0 : BitVec (w+1)) --- let s : BitVec (w+w+1) := -x ++ (0 : BitVec (w+1)) --- let p : BitVec (w+w+1) := (0 : BitVec w) ++ (y : BitVec w) ++ (0 : BitVec 1) --- go a s p w --- where --- go (a s p : BitVec (w+w+1)) : Nat → BitVec w --- | 0 => p --- | n+1 => - end BitVec diff --git a/src/Init/Data/BitVec/Lemmas.lean b/src/Init/Data/BitVec/Lemmas.lean index 13f4bf8a8a95..7c151fd24536 100644 --- a/src/Init/Data/BitVec/Lemmas.lean +++ b/src/Init/Data/BitVec/Lemmas.lean @@ -704,7 +704,7 @@ theorem msb_append {x : BitVec w} {y : BitVec v} : simp only [getLsb_append, cond_eq_if] split <;> simp [*] -theorem BitVec.shiftRight_shiftRight (w : Nat) (x : BitVec w) (n m : Nat) : +theorem shiftRight_shiftRight (w : Nat) (x : BitVec w) (n m : Nat) : (x >>> n) >>> m = x >>> (n + m) := by ext i simp [Nat.add_assoc n m i]